| author | wenzelm | 
| Fri, 05 Aug 2022 13:23:52 +0200 | |
| changeset 75760 | f8be63d2ec6f | 
| parent 74598 | 5d91897a8e54 | 
| child 76822 | 25c0d4c0e110 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Measure_Space.thy | 
| 47694 | 2 | Author: Lawrence C Paulson | 
| 3 | Author: Johannes Hölzl, TU München | |
| 4 | Author: Armin Heller, TU München | |
| 5 | *) | |
| 6 | ||
| 69517 | 7 | section \<open>Measure Spaces\<close> | 
| 47694 | 8 | |
| 9 | theory Measure_Space | |
| 10 | imports | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
64911diff
changeset | 11 | Measurable "HOL-Library.Extended_Nonnegative_Real" | 
| 47694 | 12 | begin | 
| 13 | ||
| 70136 | 14 | subsection\<^marker>\<open>tag unimportant\<close> "Relate extended reals and the indicator function" | 
| 50104 | 15 | |
| 47694 | 16 | lemma suminf_cmult_indicator: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 17 | fixes f :: "nat \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 18 | assumes "disjoint_family A" "x \<in> A i" | 
| 47694 | 19 | shows "(\<Sum>n. f n * indicator (A n) x) = f i" | 
| 20 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 21 | have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)" | 
| 61808 | 22 | using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 23 | then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)" | 
| 64267 | 24 | by (auto simp: sum.If_cases) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 25 | moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)" | 
| 51000 | 26 | proof (rule SUP_eqI) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 27 | fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y" | 
| 47694 | 28 | from this[of "Suc i"] show "f i \<le> y" by auto | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 29 | qed (insert assms, simp) | 
| 47694 | 30 | ultimately show ?thesis using assms | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 31 | by (subst suminf_eq_SUP) (auto simp: indicator_def) | 
| 47694 | 32 | qed | 
| 33 | ||
| 34 | lemma suminf_indicator: | |
| 35 | assumes "disjoint_family A" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 36 | shows "(\<Sum>n. indicator (A n) x :: ennreal) = indicator (\<Union>i. A i) x" | 
| 47694 | 37 | proof cases | 
| 38 | assume *: "x \<in> (\<Union>i. A i)" | |
| 39 | then obtain i where "x \<in> A i" by auto | |
| 61808 | 40 | from suminf_cmult_indicator[OF assms(1), OF \<open>x \<in> A i\<close>, of "\<lambda>k. 1"] | 
| 47694 | 41 | show ?thesis using * by simp | 
| 42 | qed simp | |
| 43 | ||
| 64267 | 44 | lemma sum_indicator_disjoint_family: | 
| 60727 | 45 | fixes f :: "'d \<Rightarrow> 'e::semiring_1" | 
| 46 | assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P" | |
| 47 | shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j" | |
| 48 | proof - | |
| 49 |   have "P \<inter> {i. x \<in> A i} = {j}"
 | |
| 61808 | 50 | using d \<open>x \<in> A j\<close> \<open>j \<in> P\<close> unfolding disjoint_family_on_def | 
| 60727 | 51 | by auto | 
| 73536 | 52 | with \<open>finite P\<close> show ?thesis | 
| 53 | by (simp add: indicator_def) | |
| 60727 | 54 | qed | 
| 55 | ||
| 61808 | 56 | text \<open> | 
| 69597 | 57 | The type for emeasure spaces is already defined in \<^theory>\<open>HOL-Analysis.Sigma_Algebra\<close>, as it | 
| 68484 | 58 | is also used to represent sigma algebras (with an arbitrary emeasure). | 
| 61808 | 59 | \<close> | 
| 47694 | 60 | |
| 70136 | 61 | subsection\<^marker>\<open>tag unimportant\<close> "Extend binary sets" | 
| 47694 | 62 | |
| 63 | lemma LIMSEQ_binaryset: | |
| 64 |   assumes f: "f {} = 0"
 | |
| 61969 | 65 | shows "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B" | 
| 47694 | 66 | proof - | 
| 67 | have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)" | |
| 68 | proof | |
| 69 | fix n | |
| 70 | show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B" | |
| 71 | by (induct n) (auto simp add: binaryset_def f) | |
| 72 | qed | |
| 73 | moreover | |
| 61969 | 74 | have "... \<longlonglongrightarrow> f A + f B" by (rule tendsto_const) | 
| 47694 | 75 | ultimately | 
| 61969 | 76 | have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) \<longlonglongrightarrow> f A + f B" | 
| 47694 | 77 | by metis | 
| 61969 | 78 | hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B" | 
| 47694 | 79 | by simp | 
| 80 | thus ?thesis by (rule LIMSEQ_offset [where k=2]) | |
| 81 | qed | |
| 82 | ||
| 83 | lemma binaryset_sums: | |
| 84 |   assumes f: "f {} = 0"
 | |
| 85 | shows "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)" | |
| 86 | by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) | |
| 87 | ||
| 88 | lemma suminf_binaryset_eq: | |
| 89 |   fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
 | |
| 90 |   shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
 | |
| 91 | by (metis binaryset_sums sums_unique) | |
| 92 | ||
| 70136 | 93 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of a premeasure \<^term>\<open>\<mu>\<close>\<close> | 
| 47694 | 94 | |
| 61808 | 95 | text \<open> | 
| 69597 | 96 | The definitions for \<^const>\<open>positive\<close> and \<^const>\<open>countably_additive\<close> should be here, by they are | 
| 97 | necessary to define \<^typ>\<open>'a measure\<close> in \<^theory>\<open>HOL-Analysis.Sigma_Algebra\<close>. | |
| 61808 | 98 | \<close> | 
| 47694 | 99 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 100 | definition subadditive where | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 101 |   "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 102 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 103 | lemma subadditiveD: "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"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 104 | by (auto simp add: subadditive_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 105 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 106 | definition countably_subadditive where | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 107 | "countably_subadditive M f \<longleftrightarrow> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 108 | (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 109 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 110 | lemma (in ring_of_sets) countably_subadditive_subadditive: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 111 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 112 | assumes f: "positive M f" and cs: "countably_subadditive M f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 113 | shows "subadditive M f" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 114 | proof (auto simp add: subadditive_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 115 | fix x y | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 116 |   assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 117 | hence "disjoint_family (binaryset x y)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 118 | by (auto simp add: disjoint_family_on_def binaryset_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 119 | hence "range (binaryset x y) \<subseteq> M \<longrightarrow> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 120 | (\<Union>i. binaryset x y i) \<in> M \<longrightarrow> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 121 | f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 122 | using cs by (auto simp add: countably_subadditive_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 123 |   hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 124 | f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 125 | by (simp add: range_binaryset_eq UN_binaryset_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 126 | thus "f (x \<union> y) \<le> f x + f y" using f x y | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 127 | by (auto simp add: Un o_def suminf_binaryset_eq positive_def) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 128 | qed | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 129 | |
| 47694 | 130 | definition additive where | 
| 131 |   "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
 | |
| 132 | ||
| 133 | definition increasing where | |
| 134 | "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)" | |
| 135 | ||
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 136 | lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
 | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 137 | |
| 47694 | 138 | lemma positiveD_empty: | 
| 139 |   "positive M f \<Longrightarrow> f {} = 0"
 | |
| 140 | by (auto simp add: positive_def) | |
| 141 | ||
| 142 | lemma additiveD: | |
| 143 |   "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
 | |
| 144 | by (auto simp add: additive_def) | |
| 145 | ||
| 146 | lemma increasingD: | |
| 147 | "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y" | |
| 148 | by (auto simp add: increasing_def) | |
| 149 | ||
| 50104 | 150 | lemma countably_additiveI[case_names countably]: | 
| 47694 | 151 | "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i)) | 
| 152 | \<Longrightarrow> countably_additive M f" | |
| 153 | by (simp add: countably_additive_def) | |
| 154 | ||
| 155 | lemma (in ring_of_sets) disjointed_additive: | |
| 156 | assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A" | |
| 157 | shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)" | |
| 158 | proof (induct n) | |
| 159 | case (Suc n) | |
| 160 | then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" | |
| 161 | by simp | |
| 162 | also have "\<dots> = f (A n \<union> disjointed A (Suc n))" | |
| 60727 | 163 | using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) | 
| 47694 | 164 | also have "A n \<union> disjointed A (Suc n) = A (Suc n)" | 
| 61808 | 165 | using \<open>incseq A\<close> by (auto dest: incseq_SucD simp: disjointed_mono) | 
| 47694 | 166 | finally show ?case . | 
| 167 | qed simp | |
| 168 | ||
| 169 | lemma (in ring_of_sets) additive_sum: | |
| 170 | fixes A:: "'i \<Rightarrow> 'a set" | |
| 171 | assumes f: "positive M f" and ad: "additive M f" and "finite S" | |
| 172 | and A: "A`S \<subseteq> M" | |
| 173 | and disj: "disjoint_family_on A S" | |
| 174 | shows "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)" | |
| 61808 | 175 | using \<open>finite S\<close> disj A | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 176 | proof induct | 
| 47694 | 177 | case empty show ?case using f by (simp add: positive_def) | 
| 178 | next | |
| 179 | case (insert s S) | |
| 180 |   then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
 | |
| 181 | by (auto simp add: disjoint_family_on_def neq_iff) | |
| 182 | moreover | |
| 183 | have "A s \<in> M" using insert by blast | |
| 184 | moreover have "(\<Union>i\<in>S. A i) \<in> M" | |
| 61808 | 185 | using insert \<open>finite S\<close> by auto | 
| 47694 | 186 | ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)" | 
| 187 | using ad UNION_in_sets A by (auto simp add: additive_def) | |
| 188 | with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] | |
| 189 | by (auto simp add: additive_def subset_insertI) | |
| 190 | qed | |
| 191 | ||
| 192 | lemma (in ring_of_sets) additive_increasing: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 193 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 47694 | 194 | assumes posf: "positive M f" and addf: "additive M f" | 
| 195 | shows "increasing M f" | |
| 196 | proof (auto simp add: increasing_def) | |
| 197 | fix x y | |
| 198 | assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y" | |
| 199 | then have "y - x \<in> M" by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 200 | then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le) | 
| 47694 | 201 | also have "... = f (x \<union> (y-x))" using addf | 
| 202 | by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) | |
| 203 | also have "... = f y" | |
| 204 | by (metis Un_Diff_cancel Un_absorb1 xy(3)) | |
| 205 | finally show "f x \<le> f y" by simp | |
| 206 | qed | |
| 207 | ||
| 50087 | 208 | lemma (in ring_of_sets) subadditive: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 209 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 210 | assumes f: "positive M f" "additive M f" and A: "A`S \<subseteq> M" and S: "finite S" | 
| 50087 | 211 | shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 212 | using S A | 
| 50087 | 213 | proof (induct S) | 
| 214 | case empty thus ?case using f by (auto simp: positive_def) | |
| 215 | next | |
| 216 | case (insert x F) | |
| 60585 | 217 | hence in_M: "A x \<in> M" "(\<Union>i\<in>F. A i) \<in> M" "(\<Union>i\<in>F. A i) - A x \<in> M" using A by force+ | 
| 218 | have subs: "(\<Union>i\<in>F. A i) - A x \<subseteq> (\<Union>i\<in>F. A i)" by auto | |
| 219 | have "(\<Union>i\<in>(insert x F). A i) = A x \<union> ((\<Union>i\<in>F. A i) - A x)" by auto | |
| 220 | hence "f (\<Union>i\<in>(insert x F). A i) = f (A x \<union> ((\<Union>i\<in>F. A i) - A x))" | |
| 50087 | 221 | by simp | 
| 60585 | 222 | also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)" | 
| 50087 | 223 | using f(2) by (rule additiveD) (insert in_M, auto) | 
| 60585 | 224 | also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)" | 
| 50087 | 225 | using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono) | 
| 226 | also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono) | |
| 60585 | 227 | finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp | 
| 50087 | 228 | qed | 
| 229 | ||
| 47694 | 230 | lemma (in ring_of_sets) countably_additive_additive: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 231 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 47694 | 232 | assumes posf: "positive M f" and ca: "countably_additive M f" | 
| 233 | shows "additive M f" | |
| 234 | proof (auto simp add: additive_def) | |
| 235 | fix x y | |
| 236 |   assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
 | |
| 237 | hence "disjoint_family (binaryset x y)" | |
| 238 | by (auto simp add: disjoint_family_on_def binaryset_def) | |
| 239 | hence "range (binaryset x y) \<subseteq> M \<longrightarrow> | |
| 240 | (\<Union>i. binaryset x y i) \<in> M \<longrightarrow> | |
| 241 | f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))" | |
| 242 | using ca | |
| 243 | by (simp add: countably_additive_def) | |
| 244 |   hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
 | |
| 245 | f (x \<union> y) = (\<Sum>n. f (binaryset x y n))" | |
| 246 | by (simp add: range_binaryset_eq UN_binaryset_eq) | |
| 247 | thus "f (x \<union> y) = f x + f y" using posf x y | |
| 248 | by (auto simp add: Un suminf_binaryset_eq positive_def) | |
| 249 | qed | |
| 250 | ||
| 251 | lemma (in algebra) increasing_additive_bound: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 252 | fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal" | 
| 47694 | 253 | assumes f: "positive M f" and ad: "additive M f" | 
| 254 | and inc: "increasing M f" | |
| 255 | and A: "range A \<subseteq> M" | |
| 256 | and disj: "disjoint_family A" | |
| 257 | shows "(\<Sum>i. f (A i)) \<le> f \<Omega>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 258 | proof (safe intro!: suminf_le_const) | 
| 47694 | 259 | fix N | 
| 260 |   note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
 | |
| 261 |   have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
 | |
| 262 | using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N) | |
| 263 | also have "... \<le> f \<Omega>" using space_closed A | |
| 264 | by (intro increasingD[OF inc] finite_UN) auto | |
| 265 | finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp | |
| 266 | qed (insert f A, auto simp: positive_def) | |
| 267 | ||
| 268 | lemma (in ring_of_sets) countably_additiveI_finite: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 269 | fixes \<mu> :: "'a set \<Rightarrow> ennreal" | 
| 47694 | 270 | assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>" | 
| 271 | shows "countably_additive M \<mu>" | |
| 272 | proof (rule countably_additiveI) | |
| 273 | fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F" | |
| 274 | ||
| 275 |   have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
 | |
| 276 |   from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
 | |
| 277 | ||
| 278 |   have inj_f: "inj_on f {i. F i \<noteq> {}}"
 | |
| 279 | proof (rule inj_onI, simp) | |
| 280 |     fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
 | |
| 281 | then have "f i \<in> F i" "f j \<in> F j" using f by force+ | |
| 282 | with disj * show "i = j" by (auto simp: disjoint_family_on_def) | |
| 283 | qed | |
| 284 | have "finite (\<Union>i. F i)" | |
| 285 | by (metis F(2) assms(1) infinite_super sets_into_space) | |
| 286 | ||
| 287 |   have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
 | |
| 61808 | 288 | by (auto simp: positiveD_empty[OF \<open>positive M \<mu>\<close>]) | 
| 47694 | 289 |   moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
 | 
| 290 | proof (rule finite_imageD) | |
| 291 |     from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
 | |
| 292 |     then show "finite (f`{i. F i \<noteq> {}})"
 | |
| 293 | by (rule finite_subset) fact | |
| 294 | qed fact | |
| 295 |   ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
 | |
| 296 | by (rule finite_subset) | |
| 297 | ||
| 298 |   have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
 | |
| 299 | using disj by (auto simp: disjoint_family_on_def) | |
| 300 | ||
| 301 | from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))" | |
| 47761 | 302 | by (rule suminf_finite) auto | 
| 47694 | 303 |   also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
 | 
| 64267 | 304 | using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto | 
| 47694 | 305 |   also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
 | 
| 61808 | 306 | using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto | 
| 47694 | 307 | also have "\<dots> = \<mu> (\<Union>i. F i)" | 
| 308 | by (rule arg_cong[where f=\<mu>]) auto | |
| 309 | finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" . | |
| 310 | qed | |
| 311 | ||
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 312 | lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 313 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 314 | assumes f: "positive M f" "additive M f" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 315 | shows "countably_additive M f \<longleftrightarrow> | 
| 61969 | 316 | (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i))" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 317 | unfolding countably_additive_def | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 318 | proof safe | 
| 69313 | 319 | assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> \<Union>(A ` UNIV) \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>(A ` UNIV))" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 320 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 321 | then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 322 | with count_sum[THEN spec, of "disjointed A"] A(3) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 323 | have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 324 | by (auto simp: UN_disjointed_eq disjoint_family_disjointed) | 
| 61969 | 325 | moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 326 | using f(1)[unfolded positive_def] dA | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 327 | by (auto intro!: summable_LIMSEQ) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 328 | from LIMSEQ_Suc[OF this] | 
| 61969 | 329 | have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))" | 
| 56193 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
56154diff
changeset | 330 | unfolding lessThan_Suc_atMost . | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 331 | moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 332 | using disjointed_additive[OF f A(1,2)] . | 
| 61969 | 333 | ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" by simp | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 334 | next | 
| 61969 | 335 | assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 336 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M" | 
| 57446 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 hoelzl parents: 
57418diff
changeset | 337 | have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto | 
| 61969 | 338 | have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 339 | proof (unfold *[symmetric], intro cont[rule_format]) | 
| 60585 | 340 | show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 341 | using A * by auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 342 | qed (force intro!: incseq_SucI) | 
| 57446 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 hoelzl parents: 
57418diff
changeset | 343 | moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 344 | using A | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 345 | by (intro additive_sum[OF f, of _ A, symmetric]) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 346 | (auto intro: disjoint_family_on_mono[where B=UNIV]) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 347 | ultimately | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 348 | have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)" | 
| 57446 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 hoelzl parents: 
57418diff
changeset | 349 | unfolding sums_def by simp | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 350 | from sums_unique[OF this] | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 351 | show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 352 | qed | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 353 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 354 | lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 355 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 356 | assumes f: "positive M f" "additive M f" | 
| 61969 | 357 | 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)) \<longlonglongrightarrow> f (\<Inter>i. A i)) | 
| 358 |      \<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)) \<longlonglongrightarrow> 0)"
 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 359 | proof safe | 
| 61969 | 360 | 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)) \<longlonglongrightarrow> f (\<Inter>i. A i))" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 361 |   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>"
 | 
| 61969 | 362 | with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0" | 
| 61808 | 363 | using \<open>positive M f\<close>[unfolded positive_def] by auto | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 364 | next | 
| 61969 | 365 |   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)) \<longlonglongrightarrow> 0"
 | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 366 | 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>" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 367 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 368 | have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 369 | using additive_increasing[OF f] unfolding increasing_def by simp | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 370 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 371 | have decseq_fA: "decseq (\<lambda>i. f (A i))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 372 | using A by (auto simp: decseq_def intro!: f_mono) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 373 | have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 374 | using A by (auto simp: decseq_def) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 375 | then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 376 | using A unfolding decseq_def by (auto intro!: f_mono Diff) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 377 | have "f (\<Inter>x. A x) \<le> f (A 0)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 378 | using A by (auto intro!: f_mono) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 379 | then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 380 | using A by (auto simp: top_unique) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 381 |   { fix i
 | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 382 | have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 383 | then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 384 | using A by (auto simp: top_unique) } | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 385 | note f_fin = this | 
| 61969 | 386 | have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 387 | proof (intro cont[rule_format, OF _ decseq _ f_fin]) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 388 |     show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
 | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 389 | using A by auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 390 | qed | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68484diff
changeset | 391 | from INF_Lim[OF decseq_f this] | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 392 | have "(INF n. f (A n - (\<Inter>i. A i))) = 0" . | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 393 | moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 394 | by auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 395 | ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 396 | using A(4) f_fin f_Int_fin | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 397 | by (subst INF_ennreal_add_const) (auto simp: decseq_f) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 398 |   moreover {
 | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 399 | fix n | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 400 | 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))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 401 | using A by (subst f(2)[THEN additiveD]) auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 402 | also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 403 | by auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 404 | finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . } | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 405 | ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 406 | by simp | 
| 51351 | 407 | with LIMSEQ_INF[OF decseq_fA] | 
| 61969 | 408 | show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i)" by simp | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 409 | qed | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 410 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 411 | lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 412 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 413 | assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>" | 
| 61969 | 414 |   assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
 | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 415 | assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M" | 
| 61969 | 416 | shows "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 417 | proof - | 
| 61969 | 418 | from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) \<longlonglongrightarrow> 0" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 419 | by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 420 | moreover | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 421 |   { fix i
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 422 | have "f ((\<Union>i. A i) - A i \<union> A i) = f ((\<Union>i. A i) - A i) + f (A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 423 | using A by (intro f(2)[THEN additiveD]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 424 | also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 425 | by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 426 | finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 427 | using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 428 | moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 429 | using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 430 | by (auto intro!: always_eventually simp: subset_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 431 | ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 432 | by (auto intro: ennreal_tendsto_const_minus) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 433 | qed | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 434 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 435 | lemma (in ring_of_sets) empty_continuous_imp_countably_additive: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 436 | fixes f :: "'a set \<Rightarrow> ennreal" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 437 | assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. f A \<noteq> \<infinity>" | 
| 61969 | 438 |   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
 | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 439 | shows "countably_additive M f" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 440 | using countably_additive_iff_continuous_from_below[OF f] | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 441 | using empty_continuous_imp_continuous_from_below[OF f fin] cont | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 442 | by blast | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 443 | |
| 70136 | 444 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Properties of \<^const>\<open>emeasure\<close>\<close> | 
| 47694 | 445 | |
| 446 | lemma emeasure_positive: "positive (sets M) (emeasure M)" | |
| 447 | by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) | |
| 448 | ||
| 449 | lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
 | |
| 450 | using emeasure_positive[of M] by (simp add: positive_def) | |
| 451 | ||
| 59000 | 452 | lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 453 |   using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])
 | 
| 59000 | 454 | |
| 47694 | 455 | lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" | 
| 456 | by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) | |
| 457 | ||
| 458 | lemma suminf_emeasure: | |
| 459 | "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 460 | using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M] | 
| 47694 | 461 | by (simp add: countably_additive_def) | 
| 462 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 463 | lemma sums_emeasure: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 464 | "disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 465 | unfolding sums_iff by (intro conjI suminf_emeasure) auto | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 466 | |
| 47694 | 467 | lemma emeasure_additive: "additive (sets M) (emeasure M)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 468 | by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive) | 
| 47694 | 469 | |
| 470 | lemma plus_emeasure: | |
| 471 |   "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
 | |
| 472 | using additiveD[OF emeasure_additive] .. | |
| 473 | ||
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 474 | lemma emeasure_Un: | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 475 | "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 476 | using plus_emeasure[of A M "B - A"] by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 477 | |
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 478 | lemma emeasure_Un_Int: | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 479 | assumes "A \<in> sets M" "B \<in> sets M" | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 480 | shows "emeasure M A + emeasure M B = emeasure M (A \<union> B) + emeasure M (A \<inter> B)" | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 481 | proof - | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 482 | have "A = (A-B) \<union> (A \<inter> B)" by auto | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 483 | then have "emeasure M A = emeasure M (A-B) + emeasure M (A \<inter> B)" | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 484 | by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff) | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 485 | moreover have "A \<union> B = (A-B) \<union> B" by auto | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 486 | then have "emeasure M (A \<union> B) = emeasure M (A-B) + emeasure M B" | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 487 | by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff) | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 488 | ultimately show ?thesis by (metis add.assoc add.commute) | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 489 | qed | 
| 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70722diff
changeset | 490 | |
| 64267 | 491 | lemma sum_emeasure: | 
| 47694 | 492 | "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow> | 
| 493 | (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 494 | by (metis sets.additive_sum emeasure_positive emeasure_additive) | 
| 47694 | 495 | |
| 496 | lemma emeasure_mono: | |
| 497 | "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 498 | by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD) | 
| 47694 | 499 | |
| 500 | lemma emeasure_space: | |
| 501 | "emeasure M A \<le> emeasure M (space M)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 502 | by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le) | 
| 47694 | 503 | |
| 504 | lemma emeasure_Diff: | |
| 505 | assumes finite: "emeasure M B \<noteq> \<infinity>" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 506 | and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A" | 
| 47694 | 507 | shows "emeasure M (A - B) = emeasure M A - emeasure M B" | 
| 508 | proof - | |
| 61808 | 509 | have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto | 
| 47694 | 510 | then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp | 
| 511 | also have "\<dots> = emeasure M (A - B) + emeasure M B" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 512 | by (subst plus_emeasure[symmetric]) auto | 
| 47694 | 513 | finally show "emeasure M (A - B) = emeasure M A - emeasure M B" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 514 | using finite by simp | 
| 47694 | 515 | qed | 
| 516 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 517 | lemma emeasure_compl: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 518 | "s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 519 | by (rule emeasure_Diff) (auto dest: sets.sets_into_space) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 520 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 521 | lemma Lim_emeasure_incseq: | 
| 61969 | 522 | "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) \<longlonglongrightarrow> emeasure M (\<Union>i. A i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 523 | using emeasure_countably_additive | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 524 | by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive | 
| 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 525 | emeasure_additive) | 
| 47694 | 526 | |
| 527 | lemma incseq_emeasure: | |
| 528 | assumes "range B \<subseteq> sets M" "incseq B" | |
| 529 | shows "incseq (\<lambda>i. emeasure M (B i))" | |
| 530 | using assms by (auto simp: incseq_def intro!: emeasure_mono) | |
| 531 | ||
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 532 | lemma SUP_emeasure_incseq: | 
| 47694 | 533 | assumes A: "range A \<subseteq> sets M" "incseq A" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 534 | shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)" | 
| 51000 | 535 | using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 536 | by (simp add: LIMSEQ_unique) | 
| 47694 | 537 | |
| 538 | lemma decseq_emeasure: | |
| 539 | assumes "range B \<subseteq> sets M" "decseq B" | |
| 540 | shows "decseq (\<lambda>i. emeasure M (B i))" | |
| 541 | using assms by (auto simp: decseq_def intro!: emeasure_mono) | |
| 542 | ||
| 543 | lemma INF_emeasure_decseq: | |
| 544 | assumes A: "range A \<subseteq> sets M" and "decseq A" | |
| 545 | and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | |
| 546 | shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)" | |
| 547 | proof - | |
| 548 | have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)" | |
| 549 | using A by (auto intro!: emeasure_mono) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 550 | hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by (auto simp: top_unique) | 
| 47694 | 551 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 552 | have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 553 | by (simp add: ennreal_INF_const_minus) | 
| 47694 | 554 | also have "\<dots> = (SUP n. emeasure M (A 0 - A n))" | 
| 61808 | 555 | using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto | 
| 47694 | 556 | also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)" | 
| 557 | proof (rule SUP_emeasure_incseq) | |
| 558 | show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M" | |
| 559 | using A by auto | |
| 560 | show "incseq (\<lambda>n. A 0 - A n)" | |
| 61808 | 561 | using \<open>decseq A\<close> by (auto simp add: incseq_def decseq_def) | 
| 47694 | 562 | qed | 
| 563 | also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)" | |
| 564 | using A finite * by (simp, subst emeasure_Diff) auto | |
| 565 | finally show ?thesis | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 566 | by (rule ennreal_minus_cancel[rotated 3]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 567 | (insert finite A, auto intro: INF_lower emeasure_mono) | 
| 47694 | 568 | qed | 
| 569 | ||
| 63940 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 570 | lemma INF_emeasure_decseq': | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 571 | assumes A: "\<And>i. A i \<in> sets M" and "decseq A" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 572 | and finite: "\<exists>i. emeasure M (A i) \<noteq> \<infinity>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 573 | shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 574 | proof - | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 575 | from finite obtain i where i: "emeasure M (A i) < \<infinity>" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 576 | by (auto simp: less_top) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 577 | have fin: "i \<le> j \<Longrightarrow> emeasure M (A j) < \<infinity>" for j | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 578 | by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \<open>decseq A\<close>] A) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 579 | |
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 580 | have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 581 | proof (rule INF_eq) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 582 | show "\<exists>j\<in>UNIV. emeasure M (A (j + i)) \<le> emeasure M (A i')" for i' | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 583 | by (intro bexI[of _ i'] emeasure_mono decseqD[OF \<open>decseq A\<close>] A) auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 584 | qed auto | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 585 | also have "\<dots> = emeasure M (INF n. (A (n + i)))" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 586 | using A \<open>decseq A\<close> fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 587 | also have "(INF n. (A (n + i))) = (INF n. A n)" | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 588 | by (meson INF_eq UNIV_I assms(2) decseqD le_add1) | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 589 | finally show ?thesis . | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 590 | qed | 
| 
0d82c4c94014
prove HK-integrable implies Lebesgue measurable; prove HK-integral equals Lebesgue integral for nonneg functions
 hoelzl parents: 
63658diff
changeset | 591 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 592 | lemma emeasure_INT_decseq_subset: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 593 | fixes F :: "nat \<Rightarrow> 'a set" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 594 |   assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 595 | assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 596 | and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 597 | shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i\<in>I. emeasure M (F i))" | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 598 | proof cases | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 599 | assume "finite I" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 600 | have "(\<Inter>i\<in>I. F i) = F (Max I)" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 601 | using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 602 | moreover have "(INF i\<in>I. emeasure M (F i)) = emeasure M (F (Max I))" | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 603 | using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 604 | ultimately show ?thesis | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 605 | by simp | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 606 | next | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 607 | assume "infinite I" | 
| 63040 | 608 | define L where "L n = (LEAST i. i \<in> I \<and> i \<ge> n)" for n | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 609 | have L: "L n \<in> I \<and> n \<le> L n" for n | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 610 | unfolding L_def | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 611 | proof (rule LeastI_ex) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 612 | show "\<exists>x. x \<in> I \<and> n \<le> x" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 613 |       using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
 | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 614 | by (rule_tac ccontr) (auto simp: not_le) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 615 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 616 | have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 617 | unfolding L_def by (intro Least_equality) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 618 | have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 619 | using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 620 | |
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 621 | have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 622 | proof (intro INF_emeasure_decseq[symmetric]) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 623 | show "decseq (\<lambda>i. F (L i))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 624 | using L by (intro antimonoI F L_mono) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 625 | qed (insert L fin, auto) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 626 | also have "\<dots> = (INF i\<in>I. emeasure M (F i))" | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 627 | proof (intro antisym INF_greatest) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 628 | show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 629 | by (intro INF_lower2[of i]) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 630 | qed (insert L, auto intro: INF_lower) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 631 | also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 632 | proof (intro antisym INF_greatest) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 633 | show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 634 | by (intro INF_lower2[of i]) auto | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 635 | qed (insert L, auto) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 636 | finally show ?thesis . | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 637 | qed | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61166diff
changeset | 638 | |
| 47694 | 639 | lemma Lim_emeasure_decseq: | 
| 640 | assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | |
| 61969 | 641 | shows "(\<lambda>i. emeasure M (A i)) \<longlonglongrightarrow> emeasure M (\<Inter>i. A i)" | 
| 51351 | 642 | using LIMSEQ_INF[OF decseq_emeasure, OF A] | 
| 47694 | 643 | using INF_emeasure_decseq[OF A fin] by simp | 
| 644 | ||
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 645 | lemma emeasure_lfp'[consumes 1, case_names cont measurable]: | 
| 59000 | 646 | assumes "P M" | 
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60142diff
changeset | 647 | assumes cont: "sup_continuous F" | 
| 59000 | 648 | assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)" | 
| 649 |   shows "emeasure M {x\<in>space M. lfp F x} = (SUP i. emeasure M {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
 | |
| 650 | proof - | |
| 651 |   have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
 | |
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60142diff
changeset | 652 | using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) | 
| 61808 | 653 |   moreover { fix i from \<open>P M\<close> have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
 | 
| 59000 | 654 | by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } | 
| 655 |   moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
 | |
| 656 | proof (rule incseq_SucI) | |
| 657 | fix i | |
| 658 | have "(F ^^ i) (\<lambda>x. False) \<le> (F ^^ (Suc i)) (\<lambda>x. False)" | |
| 659 | proof (induct i) | |
| 660 | case 0 show ?case by (simp add: le_fun_def) | |
| 661 | next | |
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60142diff
changeset | 662 | case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto | 
| 59000 | 663 | qed | 
| 664 |     then show "{x \<in> space M. (F ^^ i) (\<lambda>x. False) x} \<subseteq> {x \<in> space M. (F ^^ Suc i) (\<lambda>x. False) x}"
 | |
| 665 | by auto | |
| 666 | qed | |
| 667 | ultimately show ?thesis | |
| 668 | by (subst SUP_emeasure_incseq) auto | |
| 669 | qed | |
| 670 | ||
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 671 | lemma emeasure_lfp: | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 672 | assumes [simp]: "\<And>s. sets (M s) = sets N" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 673 | assumes cont: "sup_continuous F" "sup_continuous f" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 674 | assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)" | 
| 60714 
ff8aa76d6d1c
stronger induction assumption in lfp_transfer and emeasure_lfp
 hoelzl parents: 
60636diff
changeset | 675 |   assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
 | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 676 |   shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 677 | proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 678 | fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 679 |   then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 680 | unfolding SUP_apply[abs_def] | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 681 | by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 682 | qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont) | 
| 47694 | 683 | |
| 684 | lemma emeasure_subadditive_finite: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 685 | "finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 686 | by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 687 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 688 | lemma emeasure_subadditive: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 689 | "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 690 |   using emeasure_subadditive_finite[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B" M] by simp
 | 
| 47694 | 691 | |
| 692 | lemma emeasure_subadditive_countably: | |
| 693 | assumes "range f \<subseteq> sets M" | |
| 694 | shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))" | |
| 695 | proof - | |
| 696 | have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)" | |
| 697 | unfolding UN_disjointed_eq .. | |
| 698 | also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 699 | using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] | 
| 47694 | 700 | by (simp add: disjoint_family_disjointed comp_def) | 
| 701 | also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 702 | using sets.range_disjointed_sets[OF assms] assms | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 703 | by (auto intro!: suminf_le emeasure_mono disjointed_subset) | 
| 47694 | 704 | finally show ?thesis . | 
| 705 | qed | |
| 706 | ||
| 707 | lemma emeasure_insert: | |
| 708 |   assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
 | |
| 709 |   shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
 | |
| 710 | proof - | |
| 61808 | 711 |   have "{x} \<inter> A = {}" using \<open>x \<notin> A\<close> by auto
 | 
| 47694 | 712 | from plus_emeasure[OF sets this] show ?thesis by simp | 
| 713 | qed | |
| 714 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 715 | lemma emeasure_insert_ne: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 716 |   "A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
 | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 717 | by (rule emeasure_insert) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 718 | |
| 64267 | 719 | lemma emeasure_eq_sum_singleton: | 
| 47694 | 720 |   assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
| 721 |   shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
 | |
| 64267 | 722 |   using sum_emeasure[of "\<lambda>x. {x}" S M] assms
 | 
| 47694 | 723 | by (auto simp: disjoint_family_on_def subset_eq) | 
| 724 | ||
| 64267 | 725 | lemma sum_emeasure_cover: | 
| 47694 | 726 | assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M" | 
| 727 | assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)" | |
| 728 | assumes disj: "disjoint_family_on B S" | |
| 729 | shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))" | |
| 730 | proof - | |
| 731 | have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))" | |
| 64267 | 732 | proof (rule sum_emeasure) | 
| 47694 | 733 | show "disjoint_family_on (\<lambda>i. A \<inter> B i) S" | 
| 61808 | 734 | using \<open>disjoint_family_on B S\<close> | 
| 47694 | 735 | unfolding disjoint_family_on_def by auto | 
| 736 | qed (insert assms, auto) | |
| 737 | also have "(\<Union>i\<in>S. A \<inter> (B i)) = A" | |
| 738 | using A by auto | |
| 739 | finally show ?thesis by simp | |
| 740 | qed | |
| 741 | ||
| 742 | lemma emeasure_eq_0: | |
| 743 | "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 744 | by (metis emeasure_mono order_eq_iff zero_le) | 
| 47694 | 745 | |
| 746 | lemma emeasure_UN_eq_0: | |
| 747 | assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M" | |
| 60585 | 748 | shows "emeasure M (\<Union>i. N i) = 0" | 
| 47694 | 749 | proof - | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 750 | have "emeasure M (\<Union>i. N i) \<le> 0" | 
| 47694 | 751 | using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 752 | then show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 753 | by (auto intro: antisym zero_le) | 
| 47694 | 754 | qed | 
| 755 | ||
| 756 | lemma measure_eqI_finite: | |
| 757 | assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" | |
| 758 |   assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
 | |
| 759 | shows "M = N" | |
| 760 | proof (rule measure_eqI) | |
| 761 | fix X assume "X \<in> sets M" | |
| 762 | then have X: "X \<subseteq> A" by auto | |
| 763 |   then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
 | |
| 64267 | 764 | using \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) | 
| 47694 | 765 |   also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
 | 
| 64267 | 766 | using X eq by (auto intro!: sum.cong) | 
| 47694 | 767 | also have "\<dots> = emeasure N X" | 
| 64267 | 768 | using X \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) | 
| 47694 | 769 | finally show "emeasure M X = emeasure N X" . | 
| 770 | qed simp | |
| 771 | ||
| 772 | lemma measure_eqI_generator_eq: | |
| 773 | fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set" | |
| 774 | assumes "Int_stable E" "E \<subseteq> Pow \<Omega>" | |
| 775 | and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X" | |
| 776 | and M: "sets M = sigma_sets \<Omega> E" | |
| 777 | and N: "sets N = sigma_sets \<Omega> E" | |
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 778 | and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | 
| 47694 | 779 | shows "M = N" | 
| 780 | proof - | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 781 | let ?\<mu> = "emeasure M" and ?\<nu> = "emeasure N" | 
| 47694 | 782 | interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 783 | have "space M = \<Omega>" | 
| 61808 | 784 | using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \<open>sets M = sigma_sets \<Omega> E\<close> | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 785 | by blast | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 786 | |
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 787 |   { fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
 | 
| 47694 | 788 | then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto | 
| 61808 | 789 | have "?\<nu> F \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> \<open>F \<in> E\<close> eq by simp | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 790 | assume "D \<in> sets M" | 
| 61808 | 791 | with \<open>Int_stable E\<close> \<open>E \<subseteq> Pow \<Omega>\<close> have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)" | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 792 | unfolding M | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 793 | proof (induct rule: sigma_sets_induct_disjoint) | 
| 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 794 | case (basic A) | 
| 61808 | 795 | then have "F \<inter> A \<in> E" using \<open>Int_stable E\<close> \<open>F \<in> E\<close> by (auto simp: Int_stable_def) | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 796 | then show ?case using eq by auto | 
| 47694 | 797 | next | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 798 | case empty then show ?case by simp | 
| 47694 | 799 | next | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 800 | case (compl A) | 
| 47694 | 801 | then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)" | 
| 802 | and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E" | |
| 61808 | 803 | using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 804 | have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 805 | then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 806 | have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 807 | then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 808 | then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding ** | 
| 61808 | 809 | using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N) | 
| 810 | also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 811 | also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding ** | 
| 61808 | 812 | using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> \<open>?\<nu> (F \<inter> A) \<noteq> \<infinity>\<close> | 
| 47694 | 813 | by (auto intro!: emeasure_Diff[symmetric] simp: M N) | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 814 | finally show ?case | 
| 61808 | 815 | using \<open>space M = \<Omega>\<close> by auto | 
| 47694 | 816 | next | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 817 | case (union A) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 818 | then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 819 | by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 820 | with A show ?case | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 821 | by auto | 
| 49789 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 hoelzl parents: 
49784diff
changeset | 822 | qed } | 
| 47694 | 823 | note * = this | 
| 824 | show "M = N" | |
| 825 | proof (rule measure_eqI) | |
| 826 | show "sets M = sets N" | |
| 827 | using M N by simp | |
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 828 | have [simp, intro]: "\<And>i. A i \<in> sets M" | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 829 | using A(1) by (auto simp: subset_eq M) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 830 | fix F assume "F \<in> sets M" | 
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 831 | let ?D = "disjointed (\<lambda>i. F \<inter> A i)" | 
| 61808 | 832 | from \<open>space M = \<Omega>\<close> have F_eq: "F = (\<Union>i. ?D i)" | 
| 833 | using \<open>F \<in> sets M\<close>[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) | |
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 834 | have [simp, intro]: "\<And>i. ?D i \<in> sets M" | 
| 61808 | 835 | using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] \<open>F \<in> sets M\<close> | 
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 836 | by (auto simp: subset_eq) | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 837 | have "disjoint_family ?D" | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 838 | by (auto simp: disjoint_family_disjointed) | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 839 | moreover | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 840 | have "(\<Sum>i. emeasure M (?D i)) = (\<Sum>i. emeasure N (?D i))" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 841 | proof (intro arg_cong[where f=suminf] ext) | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 842 | fix i | 
| 49784 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 843 | have "A i \<inter> ?D i = ?D i" | 
| 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 hoelzl parents: 
49773diff
changeset | 844 | by (auto simp: disjointed_def) | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 845 | then show "emeasure M (?D i) = emeasure N (?D i)" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 846 | using *[of "A i" "?D i", OF _ A(3)] A(1) by auto | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 847 | qed | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 848 | ultimately show "emeasure M F = emeasure N F" | 
| 61808 | 849 | by (simp add: image_subset_iff \<open>sets M = sets N\<close>[symmetric] F_eq[symmetric] suminf_emeasure) | 
| 47694 | 850 | qed | 
| 851 | qed | |
| 852 | ||
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 853 | lemma space_empty: "space M = {} \<Longrightarrow> M = count_space {}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 854 | by (rule measure_eqI) (simp_all add: space_empty_iff) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 855 | |
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 856 | lemma measure_eqI_generator_eq_countable: | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 857 | fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 858 | assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 859 | and sets: "sets M = sigma_sets \<Omega> E" "sets N = sigma_sets \<Omega> E" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 860 | and A: "A \<subseteq> E" "(\<Union>A) = \<Omega>" "countable A" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 861 | shows "M = N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 862 | proof cases | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 863 |   assume "\<Omega> = {}"
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 864 | have *: "sigma_sets \<Omega> E = sets (sigma \<Omega> E)" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 865 | using E(2) by simp | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 866 | have "space M = \<Omega>" "space N = \<Omega>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 867 | using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 868 | then show "M = N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 869 |     unfolding \<open>\<Omega> = {}\<close> by (auto dest: space_empty)
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 870 | next | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 871 |   assume "\<Omega> \<noteq> {}" with \<open>\<Union>A = \<Omega>\<close> have "A \<noteq> {}" by auto
 | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 872 | from this \<open>countable A\<close> have rng: "range (from_nat_into A) = A" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 873 | by (rule range_from_nat_into) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 874 | show "M = N" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 875 | proof (rule measure_eqI_generator_eq[OF E sets]) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 876 | show "range (from_nat_into A) \<subseteq> E" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 877 | unfolding rng using \<open>A \<subseteq> E\<close> . | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 878 | show "(\<Union>i. from_nat_into A i) = \<Omega>" | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 879 | unfolding rng using \<open>\<Union>A = \<Omega>\<close> . | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 880 | show "emeasure M (from_nat_into A i) \<noteq> \<infinity>" for i | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 881 | using rng by (intro A) auto | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 882 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 883 | qed | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 884 | |
| 47694 | 885 | lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" | 
| 886 | proof (intro measure_eqI emeasure_measure_of_sigma) | |
| 887 | show "sigma_algebra (space M) (sets M)" .. | |
| 888 | show "positive (sets M) (emeasure M)" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 889 | by (simp add: positive_def) | 
| 47694 | 890 | show "countably_additive (sets M) (emeasure M)" | 
| 891 | by (simp add: emeasure_countably_additive) | |
| 892 | qed simp_all | |
| 893 | ||
| 61808 | 894 | subsection \<open>\<open>\<mu>\<close>-null sets\<close> | 
| 47694 | 895 | |
| 70136 | 896 | definition\<^marker>\<open>tag important\<close> null_sets :: "'a measure \<Rightarrow> 'a set set" where | 
| 47694 | 897 |   "null_sets M = {N\<in>sets M. emeasure M N = 0}"
 | 
| 898 | ||
| 899 | lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0" | |
| 900 | by (simp add: null_sets_def) | |
| 901 | ||
| 902 | lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M" | |
| 903 | unfolding null_sets_def by simp | |
| 904 | ||
| 905 | lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M" | |
| 906 | unfolding null_sets_def by simp | |
| 907 | ||
| 908 | interpretation null_sets: ring_of_sets "space M" "null_sets M" for M | |
| 47762 | 909 | proof (rule ring_of_setsI) | 
| 47694 | 910 | show "null_sets M \<subseteq> Pow (space M)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 911 | using sets.sets_into_space by auto | 
| 47694 | 912 |   show "{} \<in> null_sets M"
 | 
| 913 | by auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 914 | fix A B assume null_sets: "A \<in> null_sets M" "B \<in> null_sets M" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 915 | then have sets: "A \<in> sets M" "B \<in> sets M" | 
| 47694 | 916 | by auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 917 | then have *: "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B" | 
| 47694 | 918 | "emeasure M (A - B) \<le> emeasure M A" | 
| 919 | by (auto intro!: emeasure_subadditive emeasure_mono) | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 920 | then have "emeasure M B = 0" "emeasure M A = 0" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 921 | using null_sets by auto | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 922 | with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 923 | by (auto intro!: antisym zero_le) | 
| 47694 | 924 | qed | 
| 925 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 926 | lemma UN_from_nat_into: | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 927 |   assumes I: "countable I" "I \<noteq> {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 928 | shows "(\<Union>i\<in>I. N i) = (\<Union>i. N (from_nat_into I i))" | 
| 47694 | 929 | proof - | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 930 | have "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 931 | using I by simp | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 932 | also have "\<dots> = (\<Union>i. (N \<circ> from_nat_into I) i)" | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 933 | by simp | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 934 | finally show ?thesis by simp | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 935 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 936 | |
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 937 | lemma null_sets_UN': | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 938 | assumes "countable I" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 939 | assumes "\<And>i. i \<in> I \<Longrightarrow> N i \<in> null_sets M" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 940 | shows "(\<Union>i\<in>I. N i) \<in> null_sets M" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 941 | proof cases | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 942 |   assume "I = {}" then show ?thesis by simp
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 943 | next | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 944 |   assume "I \<noteq> {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 945 | show ?thesis | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 946 | proof (intro conjI CollectI null_setsI) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 947 | show "(\<Union>i\<in>I. N i) \<in> sets M" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 948 | using assms by (intro sets.countable_UN') auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 949 | have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))" | 
| 61808 | 950 |       unfolding UN_from_nat_into[OF \<open>countable I\<close> \<open>I \<noteq> {}\<close>]
 | 
| 951 |       using assms \<open>I \<noteq> {}\<close> by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 952 | also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)" | 
| 61808 | 953 |       using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
 | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 954 | finally show "emeasure M (\<Union>i\<in>I. N i) = 0" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 955 | by (intro antisym zero_le) simp | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 956 | qed | 
| 47694 | 957 | qed | 
| 958 | ||
| 959 | lemma null_sets_UN[intro]: | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 960 | "(\<And>i::'i::countable. N i \<in> null_sets M) \<Longrightarrow> (\<Union>i. N i) \<in> null_sets M" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 961 | by (rule null_sets_UN') auto | 
| 47694 | 962 | |
| 963 | lemma null_set_Int1: | |
| 964 | assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M" | |
| 965 | proof (intro CollectI conjI null_setsI) | |
| 966 | show "emeasure M (A \<inter> B) = 0" using assms | |
| 967 | by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto | |
| 968 | qed (insert assms, auto) | |
| 969 | ||
| 970 | lemma null_set_Int2: | |
| 971 | assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M" | |
| 972 | using assms by (subst Int_commute) (rule null_set_Int1) | |
| 973 | ||
| 974 | lemma emeasure_Diff_null_set: | |
| 975 | assumes "B \<in> null_sets M" "A \<in> sets M" | |
| 976 | shows "emeasure M (A - B) = emeasure M A" | |
| 977 | proof - | |
| 978 | have *: "A - B = (A - (A \<inter> B))" by auto | |
| 979 | have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1) | |
| 980 | then show ?thesis | |
| 981 | unfolding * using assms | |
| 982 | by (subst emeasure_Diff) auto | |
| 983 | qed | |
| 984 | ||
| 985 | lemma null_set_Diff: | |
| 986 | assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M" | |
| 987 | proof (intro CollectI conjI null_setsI) | |
| 988 | show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto | |
| 989 | qed (insert assms, auto) | |
| 990 | ||
| 991 | lemma emeasure_Un_null_set: | |
| 992 | assumes "A \<in> sets M" "B \<in> null_sets M" | |
| 993 | shows "emeasure M (A \<union> B) = emeasure M A" | |
| 994 | proof - | |
| 995 | have *: "A \<union> B = A \<union> (B - A)" by auto | |
| 996 | have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff) | |
| 997 | then show ?thesis | |
| 998 | unfolding * using assms | |
| 999 | by (subst plus_emeasure[symmetric]) auto | |
| 1000 | qed | |
| 1001 | ||
| 70722 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1002 | lemma emeasure_Un': | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1003 | assumes "A \<in> sets M" "B \<in> sets M" "A \<inter> B \<in> null_sets M" | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1004 | shows "emeasure M (A \<union> B) = emeasure M A + emeasure M B" | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1005 | proof - | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1006 | have "A \<union> B = A \<union> (B - A \<inter> B)" by blast | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1007 | also have "emeasure M \<dots> = emeasure M A + emeasure M (B - A \<inter> B)" | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1008 | using assms by (subst plus_emeasure) auto | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1009 | also have "emeasure M (B - A \<inter> B) = emeasure M B" | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1010 | using assms by (intro emeasure_Diff_null_set) auto | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1011 | finally show ?thesis . | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1012 | qed | 
| 
ae2528273eeb
A couple of new theorems, stolen from AFP entries
 paulson <lp15@cam.ac.uk> parents: 
70614diff
changeset | 1013 | |
| 61808 | 1014 | subsection \<open>The almost everywhere filter (i.e.\ quantifier)\<close> | 
| 47694 | 1015 | |
| 70136 | 1016 | definition\<^marker>\<open>tag important\<close> ae_filter :: "'a measure \<Rightarrow> 'a filter" where | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 1017 | "ae_filter M = (INF N\<in>null_sets M. principal (space M - N))" | 
| 47694 | 1018 | |
| 57276 | 1019 | abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 47694 | 1020 | "almost_everywhere M P \<equiv> eventually P (ae_filter M)" | 
| 1021 | ||
| 1022 | syntax | |
| 1023 |   "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
 | |
| 1024 | ||
| 1025 | translations | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1026 | "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)" | 
| 47694 | 1027 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1028 | abbreviation | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1029 | "set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1030 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1031 | syntax | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1032 | "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1033 |   ("AE _\<in>_ in _./ _" [0,0,0,10] 10)
 | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1034 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1035 | translations | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1036 | "AE x\<in>A in M. P" \<rightleftharpoons> "CONST set_almost_everywhere A M (\<lambda>x. P)" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1037 | |
| 57276 | 1038 | lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
 | 
| 1039 | unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq) | |
| 47694 | 1040 | |
| 1041 | lemma AE_I': | |
| 1042 |   "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
 | |
| 1043 | unfolding eventually_ae_filter by auto | |
| 1044 | ||
| 1045 | lemma AE_iff_null: | |
| 1046 |   assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
 | |
| 1047 |   shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
 | |
| 1048 | proof | |
| 1049 | assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0" | |
| 1050 | unfolding eventually_ae_filter by auto | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1051 | have "emeasure M ?P \<le> emeasure M N" | 
| 47694 | 1052 | using assms N(1,2) by (auto intro: emeasure_mono) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1053 | then have "emeasure M ?P = 0" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1054 | unfolding \<open>emeasure M N = 0\<close> by auto | 
| 47694 | 1055 | then show "?P \<in> null_sets M" using assms by auto | 
| 1056 | next | |
| 1057 | assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') | |
| 1058 | qed | |
| 1059 | ||
| 1060 | lemma AE_iff_null_sets: | |
| 1061 | "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1062 | using Int_absorb1[OF sets.sets_into_space, of N M] | 
| 47694 | 1063 | by (subst AE_iff_null) (auto simp: Int_def[symmetric]) | 
| 1064 | ||
| 47761 | 1065 | lemma AE_not_in: | 
| 1066 | "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N" | |
| 1067 | by (metis AE_iff_null_sets null_setsD2) | |
| 1068 | ||
| 47694 | 1069 | lemma AE_iff_measurable: | 
| 1070 |   "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
 | |
| 1071 | using AE_iff_null[of _ P] by auto | |
| 1072 | ||
| 1073 | lemma AE_E[consumes 1]: | |
| 1074 | assumes "AE x in M. P x" | |
| 1075 |   obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | |
| 1076 | using assms unfolding eventually_ae_filter by auto | |
| 1077 | ||
| 1078 | lemma AE_E2: | |
| 1079 |   assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
 | |
| 1080 |   shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
 | |
| 1081 | proof - | |
| 1082 |   have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
 | |
| 1083 | with AE_iff_null[of M P] assms show ?thesis by auto | |
| 1084 | qed | |
| 1085 | ||
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1086 | lemma AE_E3: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1087 | assumes "AE x in M. P x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1088 | obtains N where "\<And>x. x \<in> space M - N \<Longrightarrow> P x" "N \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1089 | using assms unfolding eventually_ae_filter by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1090 | |
| 47694 | 1091 | lemma AE_I: | 
| 1092 |   assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | |
| 1093 | shows "AE x in M. P x" | |
| 1094 | using assms unfolding eventually_ae_filter by auto | |
| 1095 | ||
| 1096 | lemma AE_mp[elim!]: | |
| 1097 | assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x" | |
| 1098 | shows "AE x in M. Q x" | |
| 1099 | proof - | |
| 1100 |   from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
 | |
| 1101 | and A: "A \<in> sets M" "emeasure M A = 0" | |
| 1102 | by (auto elim!: AE_E) | |
| 1103 | ||
| 1104 |   from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
 | |
| 1105 | and B: "B \<in> sets M" "emeasure M B = 0" | |
| 1106 | by (auto elim!: AE_E) | |
| 1107 | ||
| 1108 | show ?thesis | |
| 1109 | proof (intro AE_I) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1110 | have "emeasure M (A \<union> B) \<le> 0" | 
| 47694 | 1111 | using emeasure_subadditive[of A M B] A B by auto | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1112 | then show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1113 | using A B by auto | 
| 47694 | 1114 |     show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
 | 
| 1115 | using P imp by auto | |
| 1116 | qed | |
| 1117 | qed | |
| 1118 | ||
| 64911 | 1119 | text \<open>The next lemma is convenient to combine with a lemma whose conclusion is of the | 
| 69566 | 1120 | form \<open>AE x in M. P x = Q x\<close>: for such a lemma, there is no \<open>[symmetric]\<close> variant, | 
| 64911 | 1121 | but using \<open>AE_symmetric[OF...]\<close> will replace it.\<close> | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1122 | |
| 47694 | 1123 | (* depricated replace by laws about eventually *) | 
| 1124 | lemma | |
| 1125 | shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x" | |
| 1126 | and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x" | |
| 1127 | and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x" | |
| 1128 | and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x" | |
| 1129 | and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)" | |
| 1130 | by auto | |
| 1131 | ||
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1132 | lemma AE_symmetric: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1133 | assumes "AE x in M. P x = Q x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1134 | shows "AE x in M. Q x = P x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1135 | using assms by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1136 | |
| 47694 | 1137 | lemma AE_impI: | 
| 1138 | "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x" | |
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1139 | by fastforce | 
| 47694 | 1140 | |
| 1141 | lemma AE_measure: | |
| 1142 |   assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
 | |
| 1143 |   shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
 | |
| 1144 | proof - | |
| 74362 | 1145 | from AE_E[OF AE] obtain N | 
| 1146 |     where N: "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | |
| 1147 | by auto | |
| 47694 | 1148 | with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)" | 
| 1149 | by (intro emeasure_mono) auto | |
| 1150 | also have "\<dots> \<le> emeasure M ?P + emeasure M N" | |
| 1151 | using sets N by (intro emeasure_subadditive) auto | |
| 1152 | also have "\<dots> = emeasure M ?P" using N by simp | |
| 1153 | finally show "emeasure M ?P = emeasure M (space M)" | |
| 1154 | using emeasure_space[of M "?P"] by auto | |
| 1155 | qed | |
| 1156 | ||
| 1157 | lemma AE_space: "AE x in M. x \<in> space M" | |
| 1158 |   by (rule AE_I[where N="{}"]) auto
 | |
| 1159 | ||
| 1160 | lemma AE_I2[simp, intro]: | |
| 1161 | "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x" | |
| 1162 | using AE_space by force | |
| 1163 | ||
| 1164 | lemma AE_Ball_mp: | |
| 1165 | "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x" | |
| 1166 | by auto | |
| 1167 | ||
| 1168 | lemma AE_cong[cong]: | |
| 1169 | "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)" | |
| 1170 | by auto | |
| 1171 | ||
| 69546 
27dae626822b
prefer naming convention from datatype package for strong congruence rules
 haftmann parents: 
69541diff
changeset | 1172 | lemma AE_cong_simp: "M = N \<Longrightarrow> (\<And>x. x \<in> space N =simp=> P x = Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in N. Q x)" | 
| 64008 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 1173 | by (auto simp: simp_implies_def) | 
| 
17a20ca86d62
HOL-Probability: more about probability, prepare for Markov processes in the AFP
 hoelzl parents: 
63968diff
changeset | 1174 | |
| 47694 | 1175 | lemma AE_all_countable: | 
| 1176 | "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)" | |
| 1177 | proof | |
| 1178 | assume "\<forall>i. AE x in M. P i x" | |
| 1179 | from this[unfolded eventually_ae_filter Bex_def, THEN choice] | |
| 1180 |   obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
 | |
| 1181 |   have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
 | |
| 1182 | also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto | |
| 1183 |   finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
 | |
| 1184 | moreover from N have "(\<Union>i. N i) \<in> null_sets M" | |
| 1185 | by (intro null_sets_UN) auto | |
| 1186 | ultimately show "AE x in M. \<forall>i. P i x" | |
| 1187 | unfolding eventually_ae_filter by auto | |
| 1188 | qed auto | |
| 1189 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1190 | lemma AE_ball_countable: | 
| 59000 | 1191 | assumes [intro]: "countable X" | 
| 1192 | shows "(AE x in M. \<forall>y\<in>X. P x y) \<longleftrightarrow> (\<forall>y\<in>X. AE x in M. P x y)" | |
| 1193 | proof | |
| 1194 | assume "\<forall>y\<in>X. AE x in M. P x y" | |
| 1195 | from this[unfolded eventually_ae_filter Bex_def, THEN bchoice] | |
| 1196 |   obtain N where N: "\<And>y. y \<in> X \<Longrightarrow> N y \<in> null_sets M" "\<And>y. y \<in> X \<Longrightarrow> {x\<in>space M. \<not> P x y} \<subseteq> N y"
 | |
| 1197 | by auto | |
| 1198 |   have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. {x\<in>space M. \<not> P x y})"
 | |
| 1199 | by auto | |
| 1200 | also have "\<dots> \<subseteq> (\<Union>y\<in>X. N y)" | |
| 1201 | using N by auto | |
| 1202 |   finally have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. N y)" .
 | |
| 1203 | moreover from N have "(\<Union>y\<in>X. N y) \<in> null_sets M" | |
| 1204 | by (intro null_sets_UN') auto | |
| 1205 | ultimately show "AE x in M. \<forall>y\<in>X. P x y" | |
| 1206 | unfolding eventually_ae_filter by auto | |
| 1207 | qed auto | |
| 1208 | ||
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1209 | lemma AE_ball_countable': | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1210 | "(\<And>N. N \<in> I \<Longrightarrow> AE x in M. P N x) \<Longrightarrow> countable I \<Longrightarrow> AE x in M. \<forall>N \<in> I. P N x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1211 | unfolding AE_ball_countable by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1212 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1213 | lemma AE_pairwise: "countable F \<Longrightarrow> pairwise (\<lambda>A B. AE x in M. R x A B) F \<longleftrightarrow> (AE x in M. pairwise (R x) F)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1214 | unfolding pairwise_alt by (simp add: AE_ball_countable) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1215 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1216 | lemma AE_discrete_difference: | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1217 | assumes X: "countable X" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1218 |   assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
 | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1219 |   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1220 | shows "AE x in M. x \<notin> X" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1221 | proof - | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1222 |   have "(\<Union>x\<in>X. {x}) \<in> null_sets M"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1223 | using assms by (intro null_sets_UN') auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1224 | from AE_not_in[OF this] show "AE x in M. x \<notin> X" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1225 | by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1226 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57235diff
changeset | 1227 | |
| 47694 | 1228 | lemma AE_finite_all: | 
| 1229 | assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)" | |
| 1230 | using f by induct auto | |
| 1231 | ||
| 1232 | lemma AE_finite_allI: | |
| 1233 | assumes "finite S" | |
| 1234 | shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x" | |
| 61808 | 1235 | using AE_finite_all[OF \<open>finite S\<close>] by auto | 
| 47694 | 1236 | |
| 1237 | lemma emeasure_mono_AE: | |
| 1238 | assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" | |
| 1239 | and B: "B \<in> sets M" | |
| 1240 | shows "emeasure M A \<le> emeasure M B" | |
| 1241 | proof cases | |
| 1242 | assume A: "A \<in> sets M" | |
| 1243 |   from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
 | |
| 1244 | by (auto simp: eventually_ae_filter) | |
| 1245 | have "emeasure M A = emeasure M (A - N)" | |
| 1246 | using N A by (subst emeasure_Diff_null_set) auto | |
| 1247 | also have "emeasure M (A - N) \<le> emeasure M (B - N)" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1248 | using N A B sets.sets_into_space by (auto intro!: emeasure_mono) | 
| 47694 | 1249 | also have "emeasure M (B - N) = emeasure M B" | 
| 1250 | using N B by (subst emeasure_Diff_null_set) auto | |
| 1251 | finally show ?thesis . | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1252 | qed (simp add: emeasure_notin_sets) | 
| 47694 | 1253 | |
| 1254 | lemma emeasure_eq_AE: | |
| 1255 | assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B" | |
| 1256 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | |
| 1257 | shows "emeasure M A = emeasure M B" | |
| 1258 | using assms by (safe intro!: antisym emeasure_mono_AE) auto | |
| 1259 | ||
| 59000 | 1260 | lemma emeasure_Collect_eq_AE: | 
| 1261 | "AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> Measurable.pred M Q \<Longrightarrow> Measurable.pred M P \<Longrightarrow> | |
| 1262 |    emeasure M {x\<in>space M. P x} = emeasure M {x\<in>space M. Q x}"
 | |
| 1263 | by (intro emeasure_eq_AE) auto | |
| 1264 | ||
| 1265 | lemma emeasure_eq_0_AE: "AE x in M. \<not> P x \<Longrightarrow> emeasure M {x\<in>space M. P x} = 0"
 | |
| 1266 | using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"] | |
| 1267 |   by (cases "{x\<in>space M. P x} \<in> sets M") (simp_all add: emeasure_notin_sets)
 | |
| 1268 | ||
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1269 | lemma emeasure_0_AE: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1270 | assumes "emeasure M (space M) = 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1271 | shows "AE x in M. P x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1272 | using eventually_ae_filter assms by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1273 | |
| 60715 | 1274 | lemma emeasure_add_AE: | 
| 1275 | assumes [measurable]: "A \<in> sets M" "B \<in> sets M" "C \<in> sets M" | |
| 1276 | assumes 1: "AE x in M. x \<in> C \<longleftrightarrow> x \<in> A \<or> x \<in> B" | |
| 1277 | assumes 2: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)" | |
| 1278 | shows "emeasure M C = emeasure M A + emeasure M B" | |
| 1279 | proof - | |
| 1280 | have "emeasure M C = emeasure M (A \<union> B)" | |
| 1281 | by (rule emeasure_eq_AE) (insert 1, auto) | |
| 1282 | also have "\<dots> = emeasure M A + emeasure M (B - A)" | |
| 1283 | by (subst plus_emeasure) auto | |
| 1284 | also have "emeasure M (B - A) = emeasure M B" | |
| 1285 | by (rule emeasure_eq_AE) (insert 2, auto) | |
| 1286 | finally show ?thesis . | |
| 1287 | qed | |
| 1288 | ||
| 61808 | 1289 | subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close> | 
| 47694 | 1290 | |
| 70136 | 1291 | locale\<^marker>\<open>tag important\<close> sigma_finite_measure = | 
| 47694 | 1292 | fixes M :: "'a measure" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1293 | assumes sigma_finite_countable: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1294 | "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1295 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1296 | lemma (in sigma_finite_measure) sigma_finite: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1297 | obtains A :: "nat \<Rightarrow> 'a set" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1298 | where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1299 | proof - | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1300 | obtain A :: "'a set set" where | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1301 | [simp]: "countable A" and | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1302 | A: "A \<subseteq> sets M" "(\<Union>A) = space M" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1303 | using sigma_finite_countable by metis | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1304 | show thesis | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1305 | proof cases | 
| 61808 | 1306 |     assume "A = {}" with \<open>(\<Union>A) = space M\<close> show thesis
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1307 |       by (intro that[of "\<lambda>_. {}"]) auto
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1308 | next | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1309 |     assume "A \<noteq> {}"
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1310 | show thesis | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1311 | proof | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1312 | show "range (from_nat_into A) \<subseteq> sets M" | 
| 61808 | 1313 |         using \<open>A \<noteq> {}\<close> A by auto
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1314 | have "(\<Union>i. from_nat_into A i) = \<Union>A" | 
| 61808 | 1315 |         using range_from_nat_into[OF \<open>A \<noteq> {}\<close> \<open>countable A\<close>] by auto
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1316 | with A show "(\<Union>i. from_nat_into A i) = space M" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1317 | by auto | 
| 61808 | 1318 |     qed (intro A from_nat_into \<open>A \<noteq> {}\<close>)
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1319 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1320 | qed | 
| 47694 | 1321 | |
| 1322 | lemma (in sigma_finite_measure) sigma_finite_disjoint: | |
| 1323 | obtains A :: "nat \<Rightarrow> 'a set" | |
| 1324 | where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A" | |
| 60580 | 1325 | proof - | 
| 47694 | 1326 | obtain A :: "nat \<Rightarrow> 'a set" where | 
| 1327 | range: "range A \<subseteq> sets M" and | |
| 1328 | space: "(\<Union>i. A i) = space M" and | |
| 1329 | measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 1330 | using sigma_finite by blast | 
| 60580 | 1331 | show thesis | 
| 1332 | proof (rule that[of "disjointed A"]) | |
| 1333 | show "range (disjointed A) \<subseteq> sets M" | |
| 1334 | by (rule sets.range_disjointed_sets[OF range]) | |
| 1335 | show "(\<Union>i. disjointed A i) = space M" | |
| 1336 | and "disjoint_family (disjointed A)" | |
| 1337 | using disjoint_family_disjointed UN_disjointed_eq[of A] space range | |
| 1338 | by auto | |
| 1339 | show "emeasure M (disjointed A i) \<noteq> \<infinity>" for i | |
| 1340 | proof - | |
| 1341 | have "emeasure M (disjointed A i) \<le> emeasure M (A i)" | |
| 1342 | using range disjointed_subset[of A i] by (auto intro!: emeasure_mono) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1343 | then show ?thesis using measure[of i] by (auto simp: top_unique) | 
| 60580 | 1344 | qed | 
| 1345 | qed | |
| 47694 | 1346 | qed | 
| 1347 | ||
| 1348 | lemma (in sigma_finite_measure) sigma_finite_incseq: | |
| 1349 | obtains A :: "nat \<Rightarrow> 'a set" | |
| 1350 | where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A" | |
| 60580 | 1351 | proof - | 
| 47694 | 1352 | obtain F :: "nat \<Rightarrow> 'a set" where | 
| 1353 | F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>" | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 1354 | using sigma_finite by blast | 
| 60580 | 1355 | show thesis | 
| 1356 | proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"]) | |
| 1357 | show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M" | |
| 1358 | using F by (force simp: incseq_def) | |
| 1359 | show "(\<Union>n. \<Union>i\<le>n. F i) = space M" | |
| 1360 | proof - | |
| 1361 | from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto | |
| 1362 | with F show ?thesis by fastforce | |
| 1363 | qed | |
| 60585 | 1364 | show "emeasure M (\<Union>i\<le>n. F i) \<noteq> \<infinity>" for n | 
| 60580 | 1365 | proof - | 
| 60585 | 1366 | have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))" | 
| 60580 | 1367 | using F by (auto intro!: emeasure_subadditive_finite) | 
| 1368 | also have "\<dots> < \<infinity>" | |
| 64267 | 1369 | using F by (auto simp: sum_Pinfty less_top) | 
| 60580 | 1370 | finally show ?thesis by simp | 
| 1371 | qed | |
| 1372 | show "incseq (\<lambda>n. \<Union>i\<le>n. F i)" | |
| 1373 | by (force simp: incseq_def) | |
| 1374 | qed | |
| 47694 | 1375 | qed | 
| 1376 | ||
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1377 | lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1378 | fixes C::real | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1379 | assumes W_meas: "W \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1380 | and W_inf: "emeasure M W = \<infinity>" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1381 | obtains Z where "Z \<in> sets M" "Z \<subseteq> W" "emeasure M Z < \<infinity>" "emeasure M Z > C" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1382 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1383 | obtain A :: "nat \<Rightarrow> 'a set" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1384 | where A: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1385 | using sigma_finite_incseq by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1386 | define B where "B = (\<lambda>i. W \<inter> A i)" | 
| 64911 | 1387 | have B_meas: "\<And>i. B i \<in> sets M" using W_meas \<open>range A \<subseteq> sets M\<close> B_def by blast | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1388 | have b: "\<And>i. B i \<subseteq> W" using B_def by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1389 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1390 |   { fix i
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1391 | have "emeasure M (B i) \<le> emeasure M (A i)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1392 | using A by (intro emeasure_mono) (auto simp: B_def) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1393 | also have "emeasure M (A i) < \<infinity>" | 
| 64911 | 1394 | using \<open>\<And>i. emeasure M (A i) \<noteq> \<infinity>\<close> by (simp add: less_top) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1395 | finally have "emeasure M (B i) < \<infinity>" . } | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1396 | note c = this | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1397 | |
| 64911 | 1398 | have "W = (\<Union>i. B i)" using B_def \<open>(\<Union>i. A i) = space M\<close> W_meas by auto | 
| 1399 | moreover have "incseq B" using B_def \<open>incseq A\<close> by (simp add: incseq_def subset_eq) | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1400 | ultimately have "(\<lambda>i. emeasure M (B i)) \<longlonglongrightarrow> emeasure M W" using W_meas B_meas | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1401 | by (simp add: B_meas Lim_emeasure_incseq image_subset_iff) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1402 | then have "(\<lambda>i. emeasure M (B i)) \<longlonglongrightarrow> \<infinity>" using W_inf by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1403 | from order_tendstoD(1)[OF this, of C] | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1404 | obtain i where d: "emeasure M (B i) > C" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1405 | by (auto simp: eventually_sequentially) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1406 | have "B i \<in> sets M" "B i \<subseteq> W" "emeasure M (B i) < \<infinity>" "emeasure M (B i) > C" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1407 | using B_meas b c d by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1408 | then show ?thesis using that by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1409 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 1410 | |
| 69597 | 1411 | subsection \<open>Measure space induced by distribution of \<^const>\<open>measurable\<close>-functions\<close> | 
| 47694 | 1412 | |
| 70136 | 1413 | definition\<^marker>\<open>tag important\<close> distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
 | 
| 69564 | 1414 | "distr M N f = | 
| 1415 | measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))" | |
| 47694 | 1416 | |
| 1417 | lemma | |
| 59048 | 1418 | shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N" | 
| 47694 | 1419 | and space_distr[simp]: "space (distr M N f) = space N" | 
| 1420 | by (auto simp: distr_def) | |
| 1421 | ||
| 1422 | lemma | |
| 1423 | shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" | |
| 1424 | and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" | |
| 1425 | by (auto simp: measurable_def) | |
| 1426 | ||
| 54417 | 1427 | lemma distr_cong: | 
| 1428 | "M = K \<Longrightarrow> sets N = sets L \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> distr M N f = distr K L g" | |
| 1429 | using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong) | |
| 1430 | ||
| 47694 | 1431 | lemma emeasure_distr: | 
| 1432 | fixes f :: "'a \<Rightarrow> 'b" | |
| 1433 | assumes f: "f \<in> measurable M N" and A: "A \<in> sets N" | |
| 1434 | shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A") | |
| 1435 | unfolding distr_def | |
| 1436 | proof (rule emeasure_measure_of_sigma) | |
| 1437 | show "positive (sets N) ?\<mu>" | |
| 1438 | by (auto simp: positive_def) | |
| 1439 | ||
| 1440 | show "countably_additive (sets N) ?\<mu>" | |
| 1441 | proof (intro countably_additiveI) | |
| 1442 | fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A" | |
| 1443 | then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto | |
| 1444 | then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M" | |
| 1445 | using f by (auto simp: measurable_def) | |
| 1446 | moreover have "(\<Union>i. f -` A i \<inter> space M) \<in> sets M" | |
| 1447 | using * by blast | |
| 1448 | moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)" | |
| 61808 | 1449 | using \<open>disjoint_family A\<close> by (auto simp: disjoint_family_on_def) | 
| 47694 | 1450 | ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" | 
| 1451 | using suminf_emeasure[OF _ **] A f | |
| 1452 | by (auto simp: comp_def vimage_UN) | |
| 1453 | qed | |
| 1454 | show "sigma_algebra (space N) (sets N)" .. | |
| 1455 | qed fact | |
| 1456 | ||
| 59000 | 1457 | lemma emeasure_Collect_distr: | 
| 1458 | assumes X[measurable]: "X \<in> measurable M N" "Measurable.pred N P" | |
| 1459 |   shows "emeasure (distr M N X) {x\<in>space N. P x} = emeasure M {x\<in>space M. P (X x)}"
 | |
| 1460 | by (subst emeasure_distr) | |
| 1461 | (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space]) | |
| 1462 | ||
| 1463 | lemma emeasure_lfp2[consumes 1, case_names cont f measurable]: | |
| 1464 | assumes "P M" | |
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
60142diff
changeset | 1465 | assumes cont: "sup_continuous F" | 
| 59000 | 1466 | assumes f: "\<And>M. P M \<Longrightarrow> f \<in> measurable M' M" | 
| 1467 | assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)" | |
| 1468 |   shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
 | |
| 1469 | proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) | |
| 1470 | show "f \<in> measurable M' M" "f \<in> measurable M' M" | |
| 61808 | 1471 | using f[OF \<open>P M\<close>] by auto | 
| 59000 | 1472 |   { fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
 | 
| 61808 | 1473 | using \<open>P M\<close> by (induction i arbitrary: M) (auto intro!: *) } | 
| 59000 | 1474 | show "Measurable.pred M (lfp F)" | 
| 61808 | 1475 | using \<open>P M\<close> cont * by (rule measurable_lfp_coinduct[of P]) | 
| 59000 | 1476 | |
| 1477 |   have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
 | |
| 1478 |     (SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
 | |
| 61808 | 1479 | using \<open>P M\<close> | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 1480 | proof (coinduction arbitrary: M rule: emeasure_lfp') | 
| 59000 | 1481 | case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A" | 
| 1482 | by metis | |
| 1483 | then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A" | |
| 1484 | by simp | |
| 61808 | 1485 | with \<open>P N\<close>[THEN *] show ?case | 
| 59000 | 1486 | by auto | 
| 1487 | qed fact | |
| 1488 |   then show "emeasure (distr M' M f) {x \<in> space M. lfp F x} =
 | |
| 1489 |     (SUP i. emeasure (distr M' M f) {x \<in> space M. (F ^^ i) (\<lambda>x. False) x})"
 | |
| 1490 | by simp | |
| 1491 | qed | |
| 1492 | ||
| 50104 | 1493 | lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N" | 
| 1494 | by (rule measure_eqI) (auto simp: emeasure_distr) | |
| 1495 | ||
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 1496 | lemma distr_id2: "sets M = sets N \<Longrightarrow> distr N M (\<lambda>x. x) = N" | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 1497 | by (rule measure_eqI) (auto simp: emeasure_distr) | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 1498 | |
| 50001 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49789diff
changeset | 1499 | lemma measure_distr: | 
| 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49789diff
changeset | 1500 | "f \<in> measurable M N \<Longrightarrow> S \<in> sets N \<Longrightarrow> measure (distr M N f) S = measure M (f -` S \<inter> space M)" | 
| 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49789diff
changeset | 1501 | by (simp add: emeasure_distr measure_def) | 
| 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49789diff
changeset | 1502 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1503 | lemma distr_cong_AE: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1504 | assumes 1: "M = K" "sets N = sets L" and | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1505 | 2: "(AE x in M. f x = g x)" and "f \<in> measurable M N" and "g \<in> measurable K L" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1506 | shows "distr M N f = distr K L g" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1507 | proof (rule measure_eqI) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1508 | fix A assume "A \<in> sets (distr M N f)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1509 | with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1510 | by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1511 | qed (insert 1, simp) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 1512 | |
| 47694 | 1513 | lemma AE_distrD: | 
| 1514 | assumes f: "f \<in> measurable M M'" | |
| 1515 | and AE: "AE x in distr M M' f. P x" | |
| 1516 | shows "AE x in M. P (f x)" | |
| 1517 | proof - | |
| 74362 | 1518 | from AE[THEN AE_E] obtain N | 
| 1519 |     where "{x \<in> space (distr M M' f). \<not> P x} \<subseteq> N"
 | |
| 1520 | "emeasure (distr M M' f) N = 0" | |
| 1521 | "N \<in> sets (distr M M' f)" | |
| 1522 | by auto | |
| 47694 | 1523 | with f show ?thesis | 
| 74362 | 1524 | by (simp add: eventually_ae_filter, intro bexI[of _ "f -` N \<inter> space M"]) | 
| 47694 | 1525 | (auto simp: emeasure_distr measurable_def) | 
| 1526 | qed | |
| 1527 | ||
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 1528 | lemma AE_distr_iff: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1529 |   assumes f[measurable]: "f \<in> measurable M N" and P[measurable]: "{x \<in> space N. P x} \<in> sets N"
 | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 1530 | shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 1531 | proof (subst (1 2) AE_iff_measurable[OF _ refl]) | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1532 |   have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
 | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1533 | using f[THEN measurable_space] by auto | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1534 |   then show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
 | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 1535 |     (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
 | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1536 | by (simp add: emeasure_distr) | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1537 | qed auto | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 1538 | |
| 47694 | 1539 | lemma null_sets_distr_iff: | 
| 1540 | "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1541 | by (auto simp add: null_sets_def emeasure_distr) | 
| 47694 | 1542 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 1543 | proposition distr_distr: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1544 | "g \<in> measurable N L \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> distr (distr M N f) L g = distr M L (g \<circ> f)" | 
| 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1545 | by (auto simp add: emeasure_distr measurable_space | 
| 47694 | 1546 | intro!: arg_cong[where f="emeasure M"] measure_eqI) | 
| 1547 | ||
| 70136 | 1548 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Real measure values\<close> | 
| 47694 | 1549 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1550 | lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1551 | proof (rule ring_of_setsI) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1552 |   show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1553 |     a \<union> b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1554 | using emeasure_subadditive[of a M b] by (auto simp: top_unique) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1555 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1556 |   show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1557 |     a - b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
 | 
| 69286 | 1558 | using emeasure_mono[of "a - b" a M] by (auto simp: top_unique) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1559 | qed (auto dest: sets.sets_into_space) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1560 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1561 | lemma measure_nonneg[simp]: "0 \<le> measure M A" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 1562 | unfolding measure_def by auto | 
| 47694 | 1563 | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1564 | lemma measure_nonneg' [simp]: "\<not> measure M A < 0" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1565 | using measure_nonneg not_le by blast | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1566 | |
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1567 | lemma zero_less_measure_iff: "0 < measure M A \<longleftrightarrow> measure M A \<noteq> 0" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1568 | using measure_nonneg[of M A] by (auto simp add: le_less) | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61808diff
changeset | 1569 | |
| 59000 | 1570 | lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1571 | using measure_nonneg[of M X] by linarith | 
| 59000 | 1572 | |
| 47694 | 1573 | lemma measure_empty[simp]: "measure M {} = 0"
 | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1574 | unfolding measure_def by (simp add: zero_ennreal.rep_eq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1575 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1576 | lemma emeasure_eq_ennreal_measure: | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1577 | "emeasure M A \<noteq> top \<Longrightarrow> emeasure M A = ennreal (measure M A)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1578 | by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def) | 
| 47694 | 1579 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1580 | lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0" | 
| 71633 | 1581 | by (simp add: measure_def) | 
| 47694 | 1582 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1583 | lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1584 | using emeasure_eq_ennreal_measure[of M A] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1585 | by (cases "A \<in> M") (auto simp: measure_notin_sets emeasure_notin_sets) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1586 | |
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1587 | lemma enn2real_plus:"a < top \<Longrightarrow> b < top \<Longrightarrow> enn2real (a + b) = enn2real a + enn2real b" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 1588 | by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1589 | del: real_of_ereal_enn2ereal) | 
| 61633 | 1590 | |
| 70380 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1591 | lemma enn2real_sum:"(\<And>i. i \<in> I \<Longrightarrow> f i < top) \<Longrightarrow> enn2real (sum f I) = sum (enn2real \<circ> f) I" | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1592 | by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) | 
| 
2b0dca68c3ee
More analysis / measure theory material
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 1593 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1594 | lemma measure_eq_AE: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1595 | assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1596 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1597 | shows "measure M A = measure M B" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1598 | using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1599 | |
| 47694 | 1600 | lemma measure_Union: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1601 |   "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1602 | measure M (A \<union> B) = measure M A + measure M B" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 1603 | by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1604 | |
| 47694 | 1605 | lemma measure_finite_Union: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1606 | "finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow> | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1607 | measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1608 | by (induction S rule: finite_induct) | 
| 64267 | 1609 | (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite]) | 
| 47694 | 1610 | |
| 1611 | lemma measure_Diff: | |
| 1612 | assumes finite: "emeasure M A \<noteq> \<infinity>" | |
| 1613 | and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A" | |
| 1614 | shows "measure M (A - B) = measure M A - measure M B" | |
| 1615 | proof - | |
| 1616 | have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A" | |
| 1617 | using measurable by (auto intro!: emeasure_mono) | |
| 1618 | hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1619 | using measurable finite by (rule_tac measure_Union) (auto simp: top_unique) | 
| 61808 | 1620 | thus ?thesis using \<open>B \<subseteq> A\<close> by (auto simp: Un_absorb2) | 
| 47694 | 1621 | qed | 
| 1622 | ||
| 1623 | lemma measure_UNION: | |
| 1624 | assumes measurable: "range A \<subseteq> sets M" "disjoint_family A" | |
| 1625 | assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" | |
| 1626 | shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))" | |
| 1627 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1628 | have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1629 | unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums) | 
| 47694 | 1630 | moreover | 
| 1631 |   { fix i
 | |
| 1632 | have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)" | |
| 1633 | using measurable by (auto intro!: emeasure_mono) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1634 | then have "emeasure M (A i) = ennreal ((measure M (A i)))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1635 | using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) } | 
| 47694 | 1636 | ultimately show ?thesis using finite | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 1637 | by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all | 
| 47694 | 1638 | qed | 
| 1639 | ||
| 1640 | lemma measure_subadditive: | |
| 1641 | assumes measurable: "A \<in> sets M" "B \<in> sets M" | |
| 1642 | and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1643 | shows "measure M (A \<union> B) \<le> measure M A + measure M B" | 
| 47694 | 1644 | proof - | 
| 1645 | have "emeasure M (A \<union> B) \<noteq> \<infinity>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1646 | using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique) | 
| 47694 | 1647 | then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)" | 
| 1648 | using emeasure_subadditive[OF measurable] fin | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1649 | apply simp | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1650 | apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure) | 
| 68403 | 1651 | apply (auto simp flip: ennreal_plus) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1652 | done | 
| 47694 | 1653 | qed | 
| 1654 | ||
| 1655 | lemma measure_subadditive_finite: | |
| 1656 | assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>" | |
| 1657 | shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))" | |
| 1658 | proof - | |
| 1659 |   { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
 | |
| 1660 | using emeasure_subadditive_finite[OF A] . | |
| 1661 | also have "\<dots> < \<infinity>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1662 | using fin by (simp add: less_top A) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1663 | finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> top" by simp } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1664 | note * = this | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1665 | show ?thesis | 
| 47694 | 1666 | using emeasure_subadditive_finite[OF A] fin | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1667 | unfolding emeasure_eq_ennreal_measure[OF *] | 
| 64267 | 1668 | by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure) | 
| 47694 | 1669 | qed | 
| 1670 | ||
| 1671 | lemma measure_subadditive_countably: | |
| 1672 | assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>" | |
| 1673 | shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))" | |
| 1674 | proof - | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1675 | from fin have **: "\<And>i. emeasure M (A i) \<noteq> top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1676 | using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top) | 
| 47694 | 1677 |   { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
 | 
| 1678 | using emeasure_subadditive_countably[OF A] . | |
| 1679 | also have "\<dots> < \<infinity>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1680 | using fin by (simp add: less_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1681 | finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp } | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1682 | then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1683 | by (rule emeasure_eq_ennreal_measure[symmetric]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1684 | also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1685 | using emeasure_subadditive_countably[OF A] . | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1686 | also have "\<dots> = ennreal (\<Sum>i. measure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1687 | using fin unfolding emeasure_eq_ennreal_measure[OF **] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1688 | by (subst suminf_ennreal) (auto simp: **) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1689 | finally show ?thesis | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1690 | apply (rule ennreal_le_iff[THEN iffD1, rotated]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1691 | apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1692 | using fin | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1693 | apply (simp add: emeasure_eq_ennreal_measure[OF **]) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1694 | done | 
| 47694 | 1695 | qed | 
| 1696 | ||
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1697 | lemma measure_Un_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A \<union> B) = measure M A" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1698 | by (simp add: measure_def emeasure_Un_null_set) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1699 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1700 | lemma measure_Diff_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A - B) = measure M A" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1701 | by (simp add: measure_def emeasure_Diff_null_set) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1702 | |
| 64267 | 1703 | lemma measure_eq_sum_singleton: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1704 |   "finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
 | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1705 |     measure M S = (\<Sum>x\<in>S. measure M {x})"
 | 
| 64267 | 1706 | using emeasure_eq_sum_singleton[of S M] | 
| 1707 | by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure) | |
| 47694 | 1708 | |
| 1709 | lemma Lim_measure_incseq: | |
| 1710 | assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1711 | shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1712 | proof (rule tendsto_ennrealD) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1713 | have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1714 | using fin by (auto simp: emeasure_eq_ennreal_measure) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1715 | moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1716 | using assms emeasure_mono[of "A _" "\<Union>i. A i" M] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1717 | by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans) | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1718 | ultimately show "(\<lambda>x. ennreal (measure M (A x))) \<longlonglongrightarrow> ennreal (measure M (\<Union>i. A i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1719 | using A by (auto intro!: Lim_emeasure_incseq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1720 | qed auto | 
| 47694 | 1721 | |
| 1722 | lemma Lim_measure_decseq: | |
| 1723 | assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" | |
| 61969 | 1724 | shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1725 | proof (rule tendsto_ennrealD) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1726 | have "ennreal (measure M (\<Inter>i. A i)) = emeasure M (\<Inter>i. A i)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1727 | using fin[of 0] A emeasure_mono[of "\<Inter>i. A i" "A 0" M] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1728 | by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1729 | moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1730 | using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1731 | ultimately show "(\<lambda>x. ennreal (measure M (A x))) \<longlonglongrightarrow> ennreal (measure M (\<Inter>i. A i))" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1732 | using fin A by (auto intro!: Lim_emeasure_decseq) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 1733 | qed auto | 
| 47694 | 1734 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1735 | subsection \<open>Set of measurable sets with finite measure\<close> | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1736 | |
| 70136 | 1737 | definition\<^marker>\<open>tag important\<close> fmeasurable :: "'a measure \<Rightarrow> 'a set set" where | 
| 69564 | 1738 | "fmeasurable M = {A\<in>sets M. emeasure M A < \<infinity>}"
 | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1739 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1740 | lemma fmeasurableD[dest, measurable_dest]: "A \<in> fmeasurable M \<Longrightarrow> A \<in> sets M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1741 | by (auto simp: fmeasurable_def) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1742 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1743 | lemma fmeasurableD2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A \<noteq> top" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1744 | by (auto simp: fmeasurable_def) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1745 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1746 | lemma fmeasurableI: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> A \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1747 | by (auto simp: fmeasurable_def) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1748 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1749 | lemma fmeasurableI_null_sets: "A \<in> null_sets M \<Longrightarrow> A \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1750 | by (auto simp: fmeasurable_def) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1751 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1752 | lemma fmeasurableI2: "A \<in> fmeasurable M \<Longrightarrow> B \<subseteq> A \<Longrightarrow> B \<in> sets M \<Longrightarrow> B \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1753 | using emeasure_mono[of B A M] by (auto simp: fmeasurable_def) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1754 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1755 | lemma measure_mono_fmeasurable: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1756 | "A \<subseteq> B \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M A \<le> measure M B" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1757 | by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1758 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1759 | lemma emeasure_eq_measure2: "A \<in> fmeasurable M \<Longrightarrow> emeasure M A = measure M A" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1760 | by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1761 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1762 | interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1763 | proof (rule ring_of_setsI) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1764 |   show "fmeasurable M \<subseteq> Pow (space M)" "{} \<in> fmeasurable M"
 | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1765 | by (auto simp: fmeasurable_def dest: sets.sets_into_space) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1766 | fix a b assume *: "a \<in> fmeasurable M" "b \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1767 | then have "emeasure M (a \<union> b) \<le> emeasure M a + emeasure M b" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1768 | by (intro emeasure_subadditive) auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1769 | also have "\<dots> < top" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1770 | using * by (auto simp: fmeasurable_def) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1771 | finally show "a \<union> b \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1772 | using * by (auto intro: fmeasurableI) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1773 | show "a - b \<in> fmeasurable M" | 
| 69286 | 1774 | using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1775 | qed | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1776 | |
| 70136 | 1777 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Measurable sets formed by unions and intersections\<close> | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1778 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1779 | lemma fmeasurable_Diff: "A \<in> fmeasurable M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A - B \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1780 | using fmeasurableI2[of A M "A - B"] by auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1781 | |
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1782 | lemma fmeasurable_Int_fmeasurable: | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1783 | "\<lbrakk>S \<in> fmeasurable M; T \<in> sets M\<rbrakk> \<Longrightarrow> (S \<inter> T) \<in> fmeasurable M" | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1784 | by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1785 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1786 | lemma fmeasurable_UN: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1787 | assumes "countable I" "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> A" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "A \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1788 | shows "(\<Union>i\<in>I. F i) \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1789 | proof (rule fmeasurableI2) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1790 | show "A \<in> fmeasurable M" "(\<Union>i\<in>I. F i) \<subseteq> A" using assms by auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1791 | show "(\<Union>i\<in>I. F i) \<in> sets M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1792 | using assms by (intro sets.countable_UN') auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1793 | qed | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1794 | |
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1795 | lemma fmeasurable_INT: | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1796 | assumes "countable I" "i \<in> I" "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M" "F i \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1797 | shows "(\<Inter>i\<in>I. F i) \<in> fmeasurable M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1798 | proof (rule fmeasurableI2) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1799 | show "F i \<in> fmeasurable M" "(\<Inter>i\<in>I. F i) \<subseteq> F i" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1800 | using assms by auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1801 | show "(\<Inter>i\<in>I. F i) \<in> sets M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1802 | using assms by (intro sets.countable_INT') auto | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1803 | qed | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 1804 | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1805 | lemma measurable_measure_Diff: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1806 | assumes "A \<in> fmeasurable M" "B \<in> sets M" "B \<subseteq> A" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1807 | shows "measure M (A - B) = measure M A - measure M B" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1808 | by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1809 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1810 | lemma measurable_Un_null_set: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1811 | assumes "B \<in> null_sets M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1812 | shows "(A \<union> B \<in> fmeasurable M \<and> A \<in> sets M) \<longleftrightarrow> A \<in> fmeasurable M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1813 | using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1814 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1815 | lemma measurable_Diff_null_set: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1816 | assumes "B \<in> null_sets M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1817 | shows "(A - B) \<in> fmeasurable M \<and> A \<in> sets M \<longleftrightarrow> A \<in> fmeasurable M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1818 | using assms | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1819 | by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1820 | |
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1821 | lemma fmeasurable_Diff_D: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1822 | assumes m: "T - S \<in> fmeasurable M" "S \<in> fmeasurable M" and sub: "S \<subseteq> T" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1823 | shows "T \<in> fmeasurable M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1824 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1825 | have "T = S \<union> (T - S)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1826 | using assms by blast | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1827 | then show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1828 | by (metis m fmeasurable.Un) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1829 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1830 | |
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1831 | lemma measure_Un2: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1832 | "A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1833 | using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1834 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1835 | lemma measure_Un3: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1836 | assumes "A \<in> fmeasurable M" "B \<in> fmeasurable M" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1837 | shows "measure M (A \<union> B) = measure M A + measure M B - measure M (A \<inter> B)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1838 | proof - | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1839 | have "measure M (A \<union> B) = measure M A + measure M (B - A)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1840 | using assms by (rule measure_Un2) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1841 | also have "B - A = B - (A \<inter> B)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1842 | by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1843 | also have "measure M (B - (A \<inter> B)) = measure M B - measure M (A \<inter> B)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1844 | using assms by (intro measure_Diff) (auto simp: fmeasurable_def) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1845 | finally show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1846 | by simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1847 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1848 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1849 | lemma measure_Un_AE: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1850 | "AE x in M. x \<notin> A \<or> x \<notin> B \<Longrightarrow> A \<in> fmeasurable M \<Longrightarrow> B \<in> fmeasurable M \<Longrightarrow> | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1851 | measure M (A \<union> B) = measure M A + measure M B" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1852 | by (subst measure_Un2) (auto intro!: measure_eq_AE) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1853 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1854 | lemma measure_UNION_AE: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1855 | assumes I: "finite I" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1856 | shows "(\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. AE x in M. x \<notin> F i \<or> x \<notin> F j) I \<Longrightarrow> | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1857 | measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1858 | unfolding AE_pairwise[OF countable_finite, OF I] | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1859 | using I | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1860 | proof (induction I rule: finite_induct) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1861 | case (insert x I) | 
| 69313 | 1862 | have "measure M (F x \<union> \<Union>(F ` I)) = measure M (F x) + measure M (\<Union>(F ` I))" | 
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1863 | by (rule measure_Un_AE) (use insert in \<open>auto simp: pairwise_insert\<close>) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1864 | with insert show ?case | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1865 | by (simp add: pairwise_insert ) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 1866 | qed simp | 
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1867 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1868 | lemma measure_UNION': | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1869 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>i j. disjnt (F i) (F j)) I \<Longrightarrow> | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1870 | measure M (\<Union>i\<in>I. F i) = (\<Sum>i\<in>I. measure M (F i))" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1871 | by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1872 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1873 | lemma measure_Union_AE: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1874 | "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise (\<lambda>S T. AE x in M. x \<notin> S \<or> x \<notin> T) F \<Longrightarrow> | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1875 | measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1876 | using measure_UNION_AE[of F "\<lambda>x. x" M] by simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1877 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1878 | lemma measure_Union': | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1879 | "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> fmeasurable M) \<Longrightarrow> pairwise disjnt F \<Longrightarrow> measure M (\<Union>F) = (\<Sum>S\<in>F. measure M S)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1880 | using measure_UNION'[of F "\<lambda>x. x" M] by simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1881 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1882 | lemma measure_Un_le: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1883 | assumes "A \<in> sets M" "B \<in> sets M" shows "measure M (A \<union> B) \<le> measure M A + measure M B" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1884 | proof cases | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1885 | assume "A \<in> fmeasurable M \<and> B \<in> fmeasurable M" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1886 | with measure_subadditive[of A M B] assms show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1887 | by (auto simp: fmeasurableD2) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1888 | next | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1889 | assume "\<not> (A \<in> fmeasurable M \<and> B \<in> fmeasurable M)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1890 | then have "A \<union> B \<notin> fmeasurable M" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1891 | using fmeasurableI2[of "A \<union> B" M A] fmeasurableI2[of "A \<union> B" M B] assms by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1892 | with assms show ?thesis | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1893 | by (auto simp: fmeasurable_def measure_def less_top[symmetric]) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1894 | qed | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1895 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1896 | lemma measure_UNION_le: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1897 | "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M) \<Longrightarrow> measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1898 | proof (induction I rule: finite_induct) | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1899 | case (insert i I) | 
| 71840 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1900 | then have "measure M (\<Union>i\<in>insert i I. F i) = measure M (F i \<union> \<Union> (F ` I))" | 
| 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1901 | by simp | 
| 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1902 | also from insert have "measure M (F i \<union> \<Union> (F ` I)) \<le> measure M (F i) + measure M (\<Union> (F ` I))" | 
| 
8ed78bb0b915
Tuned some proofs in HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
71633diff
changeset | 1903 | by (intro measure_Un_le sets.finite_Union) auto | 
| 63959 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1904 | also have "measure M (\<Union>i\<in>I. F i) \<le> (\<Sum>i\<in>I. measure M (F i))" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1905 | using insert by auto | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1906 | finally show ?case | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1907 | using insert by simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1908 | qed simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1909 | |
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1910 | lemma measure_Union_le: | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1911 | "finite F \<Longrightarrow> (\<And>S. S \<in> F \<Longrightarrow> S \<in> sets M) \<Longrightarrow> measure M (\<Union>F) \<le> (\<Sum>S\<in>F. measure M S)" | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1912 | using measure_UNION_le[of F "\<lambda>x. x" M] by simp | 
| 
f77dca1abf1b
HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63958diff
changeset | 1913 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1914 | text\<open>Version for indexed union over a countable set\<close> | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1915 | lemma | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1916 | assumes "countable I" and I: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> fmeasurable M" | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1917 | and bound: "\<And>I'. I' \<subseteq> I \<Longrightarrow> finite I' \<Longrightarrow> measure M (\<Union>i\<in>I'. A i) \<le> B" | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1918 | shows fmeasurable_UN_bound: "(\<Union>i\<in>I. A i) \<in> fmeasurable M" (is ?fm) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1919 | and measure_UN_bound: "measure M (\<Union>i\<in>I. A i) \<le> B" (is ?m) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1920 | proof - | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1921 | have "B \<ge> 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1922 | using bound by force | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1923 | have "?fm \<and> ?m" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1924 | proof cases | 
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1925 |     assume "I = {}"
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1926 | with \<open>B \<ge> 0\<close> show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1927 | by simp | 
| 63968 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1928 | next | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1929 |     assume "I \<noteq> {}"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1930 | have "(\<Union>i\<in>I. A i) = (\<Union>i. (\<Union>n\<le>i. A (from_nat_into I n)))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1931 |       by (subst range_from_nat_into[symmetric, OF \<open>I \<noteq> {}\<close> \<open>countable I\<close>]) auto
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1932 | then have "emeasure M (\<Union>i\<in>I. A i) = emeasure M (\<Union>i. (\<Union>n\<le>i. A (from_nat_into I n)))" by simp | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1933 | also have "\<dots> = (SUP i. emeasure M (\<Union>n\<le>i. A (from_nat_into I n)))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1934 |       using I \<open>I \<noteq> {}\<close>[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1935 | also have "\<dots> \<le> B" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1936 | proof (intro SUP_least) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1937 | fix i :: nat | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1938 | have "emeasure M (\<Union>n\<le>i. A (from_nat_into I n)) = measure M (\<Union>n\<le>i. A (from_nat_into I n))" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1939 |         using I \<open>I \<noteq> {}\<close>[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1940 |       also have "\<dots> = measure M (\<Union>n\<in>from_nat_into I ` {..i}. A n)"
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1941 | by simp | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1942 | also have "\<dots> \<le> B" | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1943 |         by (intro ennreal_leI bound) (auto intro:  from_nat_into[OF \<open>I \<noteq> {}\<close>])
 | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1944 | finally show "emeasure M (\<Union>n\<le>i. A (from_nat_into I n)) \<le> ennreal B" . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1945 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1946 | finally have *: "emeasure M (\<Union>i\<in>I. A i) \<le> B" . | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1947 | then have ?fm | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1948 | using I \<open>countable I\<close> by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1949 | with * \<open>0\<le>B\<close> show ?thesis | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1950 | by (simp add: emeasure_eq_measure2) | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1951 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1952 | then show ?fm ?m by auto | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1953 | qed | 
| 
4359400adfe7
HOL-Analysis: the image of a negligible set under a Lipschitz continuous function is negligible (based on HOL Light proof ported by L. C. Paulson)
 hoelzl parents: 
63959diff
changeset | 1954 | |
| 67989 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1955 | text\<open>Version for big union of a countable set\<close> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1956 | lemma | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1957 | assumes "countable \<D>" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1958 | and meas: "\<And>D. D \<in> \<D> \<Longrightarrow> D \<in> fmeasurable M" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1959 | and bound: "\<And>\<E>. \<lbrakk>\<E> \<subseteq> \<D>; finite \<E>\<rbrakk> \<Longrightarrow> measure M (\<Union>\<E>) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1960 | shows fmeasurable_Union_bound: "\<Union>\<D> \<in> fmeasurable M" (is ?fm) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1961 | and measure_Union_bound: "measure M (\<Union>\<D>) \<le> B" (is ?m) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1962 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1963 | have "B \<ge> 0" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1964 | using bound by force | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1965 | have "?fm \<and> ?m" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1966 |   proof (cases "\<D> = {}")
 | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1967 | case True | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1968 | with \<open>B \<ge> 0\<close> show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1969 | by auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1970 | next | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1971 | case False | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1972 | then obtain D :: "nat \<Rightarrow> 'a set" where D: "\<D> = range D" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1973 | using \<open>countable \<D>\<close> uncountable_def by force | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1974 | have 1: "\<And>i. D i \<in> fmeasurable M" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1975 | by (simp add: D meas) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1976 | have 2: "\<And>I'. finite I' \<Longrightarrow> measure M (\<Union>x\<in>I'. D x) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1977 | by (simp add: D bound image_subset_iff) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1978 | show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1979 | unfolding D | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1980 | by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1981 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1982 | then show ?fm ?m by auto | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1983 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1984 | |
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1985 | text\<open>Version for indexed union over the type of naturals\<close> | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1986 | lemma | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1987 | fixes S :: "nat \<Rightarrow> 'a set" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1988 | assumes S: "\<And>i. S i \<in> fmeasurable M" and B: "\<And>n. measure M (\<Union>i\<le>n. S i) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1989 | shows fmeasurable_countable_Union: "(\<Union>i. S i) \<in> fmeasurable M" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1990 | and measure_countable_Union_le: "measure M (\<Union>i. S i) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1991 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1992 | have mB: "measure M (\<Union>i\<in>I. S i) \<le> B" if "finite I" for I | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1993 | proof - | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1994 | have "(\<Union>i\<in>I. S i) \<subseteq> (\<Union>i\<le>Max I. S i)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1995 | using Max_ge that by force | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1996 | then have "measure M (\<Union>i\<in>I. S i) \<le> measure M (\<Union>i \<le> Max I. S i)" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1997 | by (rule measure_mono_fmeasurable) (use S in \<open>blast+\<close>) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1998 | then show ?thesis | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1999 | using B order_trans by blast | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2000 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2001 | show "(\<Union>i. S i) \<in> fmeasurable M" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2002 | by (auto intro: fmeasurable_UN_bound [OF _ S mB]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2003 | show "measure M (\<Union>n. S n) \<le> B" | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2004 | by (auto intro: measure_UN_bound [OF _ S mB]) | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2005 | qed | 
| 
706f86afff43
more results about measure and negligibility
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 2006 | |
| 67982 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2007 | lemma measure_diff_le_measure_setdiff: | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2008 | assumes "S \<in> fmeasurable M" "T \<in> fmeasurable M" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2009 | shows "measure M S - measure M T \<le> measure M (S - T)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2010 | proof - | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2011 | have "measure M S \<le> measure M ((S - T) \<union> T)" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2012 | by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2013 | also have "\<dots> \<le> measure M (S - T) + measure M T" | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2014 | using assms by (blast intro: measure_Un_le) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2015 | finally show ?thesis | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2016 | by (simp add: algebra_simps) | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2017 | qed | 
| 
7643b005b29a
various new results on measures, integrals, etc., and some simplified proofs
 paulson <lp15@cam.ac.uk> parents: 
67962diff
changeset | 2018 | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2019 | lemma suminf_exist_split2: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2020 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2021 | assumes "summable f" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2022 | shows "(\<lambda>n. (\<Sum>k. f(k+n))) \<longlonglongrightarrow> 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2023 | by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms]) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2024 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2025 | lemma emeasure_union_summable: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2026 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2027 | and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2028 | shows "emeasure M (\<Union>n. A n) < \<infinity>" "emeasure M (\<Union>n. A n) \<le> (\<Sum>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2029 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2030 |   define B where "B = (\<lambda>N. (\<Union>n\<in>{..<N}. A n))"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2031 | have [measurable]: "B N \<in> sets M" for N unfolding B_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2032 | have "(\<lambda>N. emeasure M (B N)) \<longlonglongrightarrow> emeasure M (\<Union>N. B N)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2033 | apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2034 | moreover have "emeasure M (B N) \<le> ennreal (\<Sum>n. measure M (A n))" for N | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2035 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2036 |     have *: "(\<Sum>n\<in>{..<N}. measure M (A n)) \<le> (\<Sum>n. measure M (A n))"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2037 | using assms(3) measure_nonneg sum_le_suminf by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2038 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2039 |     have "emeasure M (B N) \<le> (\<Sum>n\<in>{..<N}. emeasure M (A n))"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2040 | unfolding B_def by (rule emeasure_subadditive_finite, auto) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2041 |     also have "... = (\<Sum>n\<in>{..<N}. ennreal(measure M (A n)))"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2042 | using assms(2) by (simp add: emeasure_eq_ennreal_measure less_top) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2043 |     also have "... = ennreal (\<Sum>n\<in>{..<N}. measure M (A n))"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2044 | by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2045 | also have "... \<le> ennreal (\<Sum>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2046 | using * by (auto simp: ennreal_leI) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2047 | finally show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2048 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2049 | ultimately have "emeasure M (\<Union>N. B N) \<le> ennreal (\<Sum>n. measure M (A n))" | 
| 68532 
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
 paulson <lp15@cam.ac.uk> parents: 
68484diff
changeset | 2050 | by (simp add: Lim_bounded) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2051 | then show "emeasure M (\<Union>n. A n) \<le> (\<Sum>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2052 | unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2053 | then show "emeasure M (\<Union>n. A n) < \<infinity>" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2054 | by (auto simp: less_top[symmetric] top_unique) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2055 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2056 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2057 | lemma borel_cantelli_limsup1: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2058 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2059 | and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2060 | shows "limsup A \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2061 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2062 | have "emeasure M (limsup A) \<le> 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2063 | proof (rule LIMSEQ_le_const) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2064 | have "(\<lambda>n. (\<Sum>k. measure M (A (k+n)))) \<longlonglongrightarrow> 0" by (rule suminf_exist_split2[OF assms(3)]) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2065 | then show "(\<lambda>n. ennreal (\<Sum>k. measure M (A (k+n)))) \<longlonglongrightarrow> 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2066 | unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2067 | have "emeasure M (limsup A) \<le> (\<Sum>k. measure M (A (k+n)))" for n | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2068 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2069 |       have I: "(\<Union>k\<in>{n..}. A k) = (\<Union>k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce)
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2070 |       have "emeasure M (limsup A) \<le> emeasure M (\<Union>k\<in>{n..}. A k)"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2071 | by (rule emeasure_mono, auto simp add: limsup_INF_SUP) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2072 | also have "... = emeasure M (\<Union>k. A (k+n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2073 | using I by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2074 | also have "... \<le> (\<Sum>k. measure M (A (k+n)))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2075 | apply (rule emeasure_union_summable) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2076 | using assms summable_ignore_initial_segment[OF assms(3), of n] by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2077 | finally show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2078 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2079 | then show "\<exists>N. \<forall>n\<ge>N. emeasure M (limsup A) \<le> (\<Sum>k. measure M (A (k+n)))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2080 | by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2081 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2082 | then show ?thesis using assms(1) measurable_limsup by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2083 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2084 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2085 | lemma borel_cantelli_AE1: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2086 | assumes [measurable]: "\<And>n. A n \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2087 | and "\<And>n. emeasure M (A n) < \<infinity>" "summable (\<lambda>n. measure M (A n))" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2088 | shows "AE x in M. eventually (\<lambda>n. x \<in> space M - A n) sequentially" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2089 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2090 | have "AE x in M. x \<notin> limsup A" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2091 | using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2092 | moreover | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2093 |   {
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2094 | fix x assume "x \<notin> limsup A" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2095 |     then obtain N where "x \<notin> (\<Union>n\<in>{N..}. A n)" unfolding limsup_INF_SUP by blast
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2096 | then have "eventually (\<lambda>n. x \<notin> A n) sequentially" using eventually_sequentially by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2097 | } | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2098 | ultimately show ?thesis by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2099 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
64267diff
changeset | 2100 | |
| 69597 | 2101 | subsection \<open>Measure spaces with \<^term>\<open>emeasure M (space M) < \<infinity>\<close>\<close> | 
| 47694 | 2102 | |
| 70136 | 2103 | locale\<^marker>\<open>tag important\<close> finite_measure = sigma_finite_measure M for M + | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2104 | assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top" | 
| 47694 | 2105 | |
| 2106 | lemma finite_measureI[Pure.intro!]: | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 2107 | "emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57446diff
changeset | 2108 |   proof qed (auto intro!: exI[of _ "{space M}"])
 | 
| 47694 | 2109 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2110 | lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2111 | using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique) | 
| 47694 | 2112 | |
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 2113 | lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M" | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 2114 | by (auto simp: fmeasurable_def less_top[symmetric]) | 
| 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 2115 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2116 | lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2117 | by (intro emeasure_eq_ennreal_measure) simp | 
| 47694 | 2118 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2119 | lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ennreal r" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2120 | using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto | 
| 47694 | 2121 | |
| 2122 | lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)" | |
| 2123 | using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) | |
| 2124 | ||
| 2125 | lemma (in finite_measure) finite_measure_Diff: | |
| 2126 | assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A" | |
| 2127 | shows "measure M (A - B) = measure M A - measure M B" | |
| 2128 | using measure_Diff[OF _ assms] by simp | |
| 2129 | ||
| 2130 | lemma (in finite_measure) finite_measure_Union: | |
| 2131 |   assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
 | |
| 2132 | shows "measure M (A \<union> B) = measure M A + measure M B" | |
| 2133 | using measure_Union[OF _ _ assms] by simp | |
| 2134 | ||
| 2135 | lemma (in finite_measure) finite_measure_finite_Union: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2136 | assumes measurable: "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S" | 
| 47694 | 2137 | shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))" | 
| 2138 | using measure_finite_Union[OF assms] by simp | |
| 2139 | ||
| 2140 | lemma (in finite_measure) finite_measure_UNION: | |
| 2141 | assumes A: "range A \<subseteq> sets M" "disjoint_family A" | |
| 2142 | shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))" | |
| 2143 | using measure_UNION[OF A] by simp | |
| 2144 | ||
| 2145 | lemma (in finite_measure) finite_measure_mono: | |
| 2146 | assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B" | |
| 2147 | using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) | |
| 2148 | ||
| 2149 | lemma (in finite_measure) finite_measure_subadditive: | |
| 2150 | assumes m: "A \<in> sets M" "B \<in> sets M" | |
| 2151 | shows "measure M (A \<union> B) \<le> measure M A + measure M B" | |
| 2152 | using measure_subadditive[OF m] by simp | |
| 2153 | ||
| 2154 | lemma (in finite_measure) finite_measure_subadditive_finite: | |
| 2155 | assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))" | |
| 2156 | using measure_subadditive_finite[OF assms] by simp | |
| 2157 | ||
| 2158 | lemma (in finite_measure) finite_measure_subadditive_countably: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2159 | "range A \<subseteq> sets M \<Longrightarrow> summable (\<lambda>i. measure M (A i)) \<Longrightarrow> measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2160 | by (rule measure_subadditive_countably) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2161 | (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure) | 
| 47694 | 2162 | |
| 64267 | 2163 | lemma (in finite_measure) finite_measure_eq_sum_singleton: | 
| 47694 | 2164 |   assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
| 2165 |   shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
 | |
| 64267 | 2166 | using measure_eq_sum_singleton[OF assms] by simp | 
| 47694 | 2167 | |
| 2168 | lemma (in finite_measure) finite_Lim_measure_incseq: | |
| 2169 | assumes A: "range A \<subseteq> sets M" "incseq A" | |
| 61969 | 2170 | shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)" | 
| 47694 | 2171 | using Lim_measure_incseq[OF A] by simp | 
| 2172 | ||
| 2173 | lemma (in finite_measure) finite_Lim_measure_decseq: | |
| 2174 | assumes A: "range A \<subseteq> sets M" "decseq A" | |
| 61969 | 2175 | shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)" | 
| 47694 | 2176 | using Lim_measure_decseq[OF A] by simp | 
| 2177 | ||
| 2178 | lemma (in finite_measure) finite_measure_compl: | |
| 2179 | assumes S: "S \<in> sets M" | |
| 2180 | shows "measure M (space M - S) = measure M (space M) - measure M S" | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2181 | using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp | 
| 47694 | 2182 | |
| 2183 | lemma (in finite_measure) finite_measure_mono_AE: | |
| 2184 | assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M" | |
| 2185 | shows "measure M A \<le> measure M B" | |
| 2186 | using assms emeasure_mono_AE[OF imp B] | |
| 2187 | by (simp add: emeasure_eq_measure) | |
| 2188 | ||
| 2189 | lemma (in finite_measure) finite_measure_eq_AE: | |
| 2190 | assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B" | |
| 2191 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | |
| 2192 | shows "measure M A = measure M B" | |
| 2193 | using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) | |
| 2194 | ||
| 50104 | 2195 | lemma (in finite_measure) measure_increasing: "increasing M (measure M)" | 
| 2196 | by (auto intro!: finite_measure_mono simp: increasing_def) | |
| 2197 | ||
| 2198 | lemma (in finite_measure) measure_zero_union: | |
| 2199 | assumes "s \<in> sets M" "t \<in> sets M" "measure M t = 0" | |
| 2200 | shows "measure M (s \<union> t) = measure M s" | |
| 2201 | using assms | |
| 2202 | proof - | |
| 2203 | have "measure M (s \<union> t) \<le> measure M s" | |
| 2204 | using finite_measure_subadditive[of s t] assms by auto | |
| 2205 | moreover have "measure M (s \<union> t) \<ge> measure M s" | |
| 2206 | using assms by (blast intro: finite_measure_mono) | |
| 2207 | ultimately show ?thesis by simp | |
| 2208 | qed | |
| 2209 | ||
| 2210 | lemma (in finite_measure) measure_eq_compl: | |
| 2211 | assumes "s \<in> sets M" "t \<in> sets M" | |
| 2212 | assumes "measure M (space M - s) = measure M (space M - t)" | |
| 2213 | shows "measure M s = measure M t" | |
| 2214 | using assms finite_measure_compl by auto | |
| 2215 | ||
| 2216 | lemma (in finite_measure) measure_eq_bigunion_image: | |
| 2217 | assumes "range f \<subseteq> sets M" "range g \<subseteq> sets M" | |
| 2218 | assumes "disjoint_family f" "disjoint_family g" | |
| 2219 | assumes "\<And> n :: nat. measure M (f n) = measure M (g n)" | |
| 60585 | 2220 | shows "measure M (\<Union>i. f i) = measure M (\<Union>i. g i)" | 
| 50104 | 2221 | using assms | 
| 2222 | proof - | |
| 60585 | 2223 | have a: "(\<lambda> i. measure M (f i)) sums (measure M (\<Union>i. f i))" | 
| 50104 | 2224 | by (rule finite_measure_UNION[OF assms(1,3)]) | 
| 60585 | 2225 | have b: "(\<lambda> i. measure M (g i)) sums (measure M (\<Union>i. g i))" | 
| 50104 | 2226 | by (rule finite_measure_UNION[OF assms(2,4)]) | 
| 2227 | show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp | |
| 2228 | qed | |
| 2229 | ||
| 2230 | lemma (in finite_measure) measure_countably_zero: | |
| 2231 | assumes "range c \<subseteq> sets M" | |
| 2232 | assumes "\<And> i. measure M (c i) = 0" | |
| 60585 | 2233 | shows "measure M (\<Union>i :: nat. c i) = 0" | 
| 50104 | 2234 | proof (rule antisym) | 
| 60585 | 2235 | show "measure M (\<Union>i :: nat. c i) \<le> 0" | 
| 50104 | 2236 | using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2237 | qed simp | 
| 50104 | 2238 | |
| 2239 | lemma (in finite_measure) measure_space_inter: | |
| 2240 | assumes events:"s \<in> sets M" "t \<in> sets M" | |
| 2241 | assumes "measure M t = measure M (space M)" | |
| 2242 | shows "measure M (s \<inter> t) = measure M s" | |
| 2243 | proof - | |
| 2244 | have "measure M ((space M - s) \<union> (space M - t)) = measure M (space M - s)" | |
| 2245 | using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union) | |
| 2246 | also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)" | |
| 2247 | by blast | |
| 2248 | finally show "measure M (s \<inter> t) = measure M s" | |
| 2249 | using events by (auto intro!: measure_eq_compl[of "s \<inter> t" s]) | |
| 2250 | qed | |
| 2251 | ||
| 2252 | lemma (in finite_measure) measure_equiprobable_finite_unions: | |
| 2253 |   assumes s: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> sets M"
 | |
| 2254 |   assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> measure M {x} = measure M {y}"
 | |
| 2255 |   shows "measure M s = real (card s) * measure M {SOME x. x \<in> s}"
 | |
| 2256 | proof cases | |
| 2257 |   assume "s \<noteq> {}"
 | |
| 2258 | then have "\<exists> x. x \<in> s" by blast | |
| 2259 | from someI_ex[OF this] assms | |
| 2260 |   have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
 | |
| 2261 |   have "measure M s = (\<Sum> x \<in> s. measure M {x})"
 | |
| 64267 | 2262 | using finite_measure_eq_sum_singleton[OF s] by simp | 
| 50104 | 2263 |   also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
 | 
| 2264 |   also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
 | |
| 64267 | 2265 | using sum_constant assms by simp | 
| 50104 | 2266 | finally show ?thesis by simp | 
| 2267 | qed simp | |
| 2268 | ||
| 2269 | lemma (in finite_measure) measure_real_sum_image_fn: | |
| 2270 | assumes "e \<in> sets M" | |
| 2271 | assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> sets M" | |
| 2272 | assumes "finite s" | |
| 2273 |   assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
 | |
| 60585 | 2274 | assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)" | 
| 50104 | 2275 | shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))" | 
| 2276 | proof - | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 2277 | have "e \<subseteq> (\<Union>i\<in>s. f i)" | 
| 61808 | 2278 | using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 2279 | then have e: "e = (\<Union>i \<in> s. e \<inter> f i)" | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 2280 | by auto | 
| 60585 | 2281 | hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp | 
| 50104 | 2282 | also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))" | 
| 2283 | proof (rule finite_measure_finite_Union) | |
| 2284 | show "finite s" by fact | |
| 2285 | show "(\<lambda>i. e \<inter> f i)`s \<subseteq> sets M" using assms(2) by auto | |
| 2286 | show "disjoint_family_on (\<lambda>i. e \<inter> f i) s" | |
| 2287 | using disjoint by (auto simp: disjoint_family_on_def) | |
| 2288 | qed | |
| 2289 | finally show ?thesis . | |
| 2290 | qed | |
| 2291 | ||
| 2292 | lemma (in finite_measure) measure_exclude: | |
| 2293 | assumes "A \<in> sets M" "B \<in> sets M" | |
| 2294 |   assumes "measure M A = measure M (space M)" "A \<inter> B = {}"
 | |
| 2295 | shows "measure M B = 0" | |
| 2296 | using measure_space_inter[of B A] assms by (auto simp: ac_simps) | |
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2297 | lemma (in finite_measure) finite_measure_distr: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2298 | assumes f: "f \<in> measurable M M'" | 
| 57235 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2299 | shows "finite_measure (distr M M' f)" | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2300 | proof (rule finite_measureI) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2301 | have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2302 | with f show "emeasure (distr M M' f) (space (distr M M' f)) \<noteq> \<infinity>" by (auto simp: emeasure_distr) | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2303 | qed | 
| 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 hoelzl parents: 
57137diff
changeset | 2304 | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2305 | lemma emeasure_gfp[consumes 1, case_names cont measurable]: | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2306 | assumes sets[simp]: "\<And>s. sets (M s) = sets N" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2307 | assumes "\<And>s. finite_measure (M s)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2308 | assumes cont: "inf_continuous F" "inf_continuous f" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2309 | assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2310 |   assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2311 | assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2312 |   shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2313 | proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2314 | P="Measurable.pred N", symmetric]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2315 | interpret finite_measure "M s" for s by fact | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2316 | fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2317 |   then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2318 | unfolding INF_apply[abs_def] | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2319 | by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2320 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2321 |   show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
 | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2322 | using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2323 | qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60585diff
changeset | 2324 | |
| 70136 | 2325 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Counting space\<close> | 
| 47694 | 2326 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2327 | lemma strict_monoI_Suc: | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2328 | assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2329 | unfolding strict_mono_def | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2330 | proof safe | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2331 | fix n m :: nat assume "n < m" then show "f n < f m" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2332 | by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2333 | qed | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2334 | |
| 47694 | 2335 | lemma emeasure_count_space: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2336 | assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)" | 
| 47694 | 2337 | (is "_ = ?M X") | 
| 2338 | unfolding count_space_def | |
| 2339 | proof (rule emeasure_measure_of_sigma) | |
| 61808 | 2340 | show "X \<in> Pow A" using \<open>X \<subseteq> A\<close> by auto | 
| 47694 | 2341 | show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2342 | show positive: "positive (Pow A) ?M" | 
| 47694 | 2343 | by (auto simp: positive_def) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2344 | have additive: "additive (Pow A) ?M" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2345 | by (auto simp: card_Un_disjoint additive_def) | 
| 47694 | 2346 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2347 | interpret ring_of_sets A "Pow A" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2348 | by (rule ring_of_setsI) auto | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2349 | show "countably_additive (Pow A) ?M" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2350 | unfolding countably_additive_iff_continuous_from_below[OF positive additive] | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2351 | proof safe | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2352 | fix F :: "nat \<Rightarrow> 'a set" assume "incseq F" | 
| 61969 | 2353 | show "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2354 | proof cases | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2355 | assume "\<exists>i. \<forall>j\<ge>i. F i = F j" | 
| 74362 | 2356 | then obtain i where i: "\<forall>j\<ge>i. F i = F j" .. | 
| 2357 | with \<open>incseq F\<close> have "F j \<subseteq> F i" for j | |
| 2358 | by (cases "i \<le> j") (auto simp: incseq_def) | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2359 | then have eq: "(\<Union>i. F i) = F i" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2360 | by auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2361 | with i show ?thesis | 
| 70532 
fcf3b891ccb1
new material; rotated premises of Lim_transform_eventually
 paulson <lp15@cam.ac.uk> parents: 
70380diff
changeset | 2362 | by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i]) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2363 | next | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2364 | assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)" | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 2365 | then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis | 
| 61808 | 2366 | then have "\<And>i. F i \<subseteq> F (f i)" using \<open>incseq F\<close> by (auto simp: incseq_def) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51351diff
changeset | 2367 | with f have *: "\<And>i. F i \<subset> F (f i)" by auto | 
| 47694 | 2368 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2369 | have "incseq (\<lambda>i. ?M (F i))" | 
| 61808 | 2370 | using \<open>incseq F\<close> unfolding incseq_def by (auto simp: card_mono dest: finite_subset) | 
| 61969 | 2371 | then have "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> (SUP n. ?M (F n))" | 
| 51000 | 2372 | by (rule LIMSEQ_SUP) | 
| 47694 | 2373 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2374 | moreover have "(SUP n. ?M (F n)) = top" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2375 | proof (rule ennreal_SUP_eq_top) | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2376 | fix n :: nat show "\<exists>k::nat\<in>UNIV. of_nat n \<le> ?M (F k)" | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2377 | proof (induct n) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2378 | case (Suc n) | 
| 74362 | 2379 | then obtain k where "of_nat n \<le> ?M (F k)" .. | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2380 | moreover have "finite (F k) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))" | 
| 61808 | 2381 | using \<open>F k \<subset> F (f k)\<close> by (simp add: psubset_card_mono) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2382 | moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)" | 
| 61808 | 2383 | using \<open>k \<le> f k\<close> \<open>incseq F\<close> by (auto simp: incseq_def dest: finite_subset) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2384 | ultimately show ?case | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2385 | by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc) | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2386 | qed auto | 
| 47694 | 2387 | qed | 
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2388 | |
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2389 | moreover | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2390 | have "inj (\<lambda>n. F ((f ^^ n) 0))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2391 | by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2392 | then have 1: "infinite (range (\<lambda>i. F ((f ^^ i) 0)))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2393 | by (rule range_inj_infinite) | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2394 | have "infinite (Pow (\<Union>i. F i))" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2395 | by (rule infinite_super[OF _ 1]) auto | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2396 | then have "infinite (\<Union>i. F i)" | 
| 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2397 | by auto | 
| 69661 | 2398 | ultimately show ?thesis by (simp only:) simp | 
| 74598 
5d91897a8e54
moved a theorem to a sensible place
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 2399 | |
| 49773 
16907431e477
tuned measurable_If; moved countably_additive equalities to Measure_Space; tuned proofs
 hoelzl parents: 
47762diff
changeset | 2400 | qed | 
| 47694 | 2401 | qed | 
| 2402 | qed | |
| 2403 | ||
| 59011 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2404 | lemma distr_bij_count_space: | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2405 | assumes f: "bij_betw f A B" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2406 | shows "distr (count_space A) (count_space B) f = count_space B" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2407 | proof (rule measure_eqI) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2408 | have f': "f \<in> measurable (count_space A) (count_space B)" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2409 | using f unfolding Pi_def bij_betw_def by auto | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2410 | fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2411 | then have X: "X \<in> sets (count_space B)" by auto | 
| 63540 | 2412 | moreover from X have "f -` X \<inter> A = the_inv_into A f ` X" | 
| 59011 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2413 | using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric]) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2414 | moreover have "inj_on (the_inv_into A f) B" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2415 | using X f by (auto simp: bij_betw_def inj_on_the_inv_into) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2416 | with X have "inj_on (the_inv_into A f) X" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2417 | by (auto intro: subset_inj_on) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2418 | ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2419 | using f unfolding emeasure_distr[OF f' X] | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2420 | by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2421 | qed simp | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59000diff
changeset | 2422 | |
| 47694 | 2423 | lemma emeasure_count_space_finite[simp]: | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2424 | "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = of_nat (card X)" | 
| 47694 | 2425 | using emeasure_count_space[of X A] by simp | 
| 2426 | ||
| 2427 | lemma emeasure_count_space_infinite[simp]: | |
| 2428 | "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>" | |
| 2429 | using emeasure_count_space[of X A] by simp | |
| 2430 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2431 | lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then of_nat (card X) else 0)" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2432 | by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2433 | measure_zero_top measure_eq_emeasure_eq_ennreal) | 
| 58606 | 2434 | |
| 47694 | 2435 | lemma emeasure_count_space_eq_0: | 
| 2436 |   "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
 | |
| 2437 | proof cases | |
| 2438 | assume X: "X \<subseteq> A" | |
| 2439 | then show ?thesis | |
| 2440 | proof (intro iffI impI) | |
| 2441 | assume "emeasure (count_space A) X = 0" | |
| 2442 |     with X show "X = {}"
 | |
| 62390 | 2443 | by (subst (asm) emeasure_count_space) (auto split: if_split_asm) | 
| 47694 | 2444 | qed simp | 
| 2445 | qed (simp add: emeasure_notin_sets) | |
| 2446 | ||
| 2447 | lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
 | |
| 2448 | unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) | |
| 2449 | ||
| 2450 | lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)" | |
| 2451 | unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) | |
| 2452 | ||
| 57025 | 2453 | lemma sigma_finite_measure_count_space_countable: | 
| 2454 | assumes A: "countable A" | |
| 47694 | 2455 | shows "sigma_finite_measure (count_space A)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2456 |   proof qed (insert A, auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"])
 | 
| 47694 | 2457 | |
| 57025 | 2458 | lemma sigma_finite_measure_count_space: | 
| 2459 | fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)" | |
| 2460 | by (rule sigma_finite_measure_count_space_countable) auto | |
| 2461 | ||
| 47694 | 2462 | lemma finite_measure_count_space: | 
| 2463 | assumes [simp]: "finite A" | |
| 2464 | shows "finite_measure (count_space A)" | |
| 2465 | by rule simp | |
| 2466 | ||
| 2467 | lemma sigma_finite_measure_count_space_finite: | |
| 2468 | assumes A: "finite A" shows "sigma_finite_measure (count_space A)" | |
| 2469 | proof - | |
| 2470 | interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) | |
| 2471 | show "sigma_finite_measure (count_space A)" .. | |
| 2472 | qed | |
| 2473 | ||
| 70136 | 2474 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Measure restricted to space\<close> | 
| 54417 | 2475 | |
| 2476 | lemma emeasure_restrict_space: | |
| 57025 | 2477 | assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>" | 
| 54417 | 2478 | shows "emeasure (restrict_space M \<Omega>) A = emeasure M A" | 
| 63540 | 2479 | proof (cases "A \<in> sets M") | 
| 2480 | case True | |
| 57025 | 2481 | show ?thesis | 
| 54417 | 2482 | proof (rule emeasure_measure_of[OF restrict_space_def]) | 
| 67399 | 2483 | show "(\<inter>) \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)" | 
| 61808 | 2484 | using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space) | 
| 57025 | 2485 | show "positive (sets (restrict_space M \<Omega>)) (emeasure M)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2486 | by (auto simp: positive_def) | 
| 57025 | 2487 | show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)" | 
| 54417 | 2488 | proof (rule countably_additiveI) | 
| 2489 | fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A" | |
| 2490 | with assms have "\<And>i. A i \<in> sets M" "\<And>i. A i \<subseteq> space M" "disjoint_family A" | |
| 57025 | 2491 | by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff | 
| 2492 | dest: sets.sets_into_space)+ | |
| 2493 | then show "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)" | |
| 54417 | 2494 | by (subst suminf_emeasure) (auto simp: disjoint_family_subset) | 
| 2495 | qed | |
| 2496 | qed | |
| 2497 | next | |
| 63540 | 2498 | case False | 
| 2499 | with assms have "A \<notin> sets (restrict_space M \<Omega>)" | |
| 54417 | 2500 | by (simp add: sets_restrict_space_iff) | 
| 63540 | 2501 | with False show ?thesis | 
| 54417 | 2502 | by (simp add: emeasure_notin_sets) | 
| 2503 | qed | |
| 2504 | ||
| 57137 | 2505 | lemma measure_restrict_space: | 
| 2506 | assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>" | |
| 2507 | shows "measure (restrict_space M \<Omega>) A = measure M A" | |
| 2508 | using emeasure_restrict_space[OF assms] by (simp add: measure_def) | |
| 2509 | ||
| 2510 | lemma AE_restrict_space_iff: | |
| 2511 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 2512 | shows "(AE x in restrict_space M \<Omega>. P x) \<longleftrightarrow> (AE x in M. x \<in> \<Omega> \<longrightarrow> P x)" | |
| 2513 | proof - | |
| 2514 | have ex_cong: "\<And>P Q f. (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> P (f x)) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> (\<exists>x. Q x)" | |
| 2515 | by auto | |
| 2516 |   { fix X assume X: "X \<in> sets M" "emeasure M X = 0"
 | |
| 2517 | then have "emeasure M (\<Omega> \<inter> space M \<inter> X) \<le> emeasure M X" | |
| 2518 | by (intro emeasure_mono) auto | |
| 2519 | then have "emeasure M (\<Omega> \<inter> space M \<inter> X) = 0" | |
| 2520 | using X by (auto intro!: antisym) } | |
| 2521 | with assms show ?thesis | |
| 2522 | unfolding eventually_ae_filter | |
| 2523 | by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff | |
| 2524 | emeasure_restrict_space cong: conj_cong | |
| 2525 | intro!: ex_cong[where f="\<lambda>X. (\<Omega> \<inter> space M) \<inter> X"]) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2526 | qed | 
| 57137 | 2527 | |
| 57025 | 2528 | lemma restrict_restrict_space: | 
| 2529 | assumes "A \<inter> space M \<in> sets M" "B \<inter> space M \<in> sets M" | |
| 2530 | shows "restrict_space (restrict_space M A) B = restrict_space M (A \<inter> B)" (is "?l = ?r") | |
| 2531 | proof (rule measure_eqI[symmetric]) | |
| 2532 | show "sets ?r = sets ?l" | |
| 2533 | unfolding sets_restrict_space image_comp by (intro image_cong) auto | |
| 2534 | next | |
| 2535 | fix X assume "X \<in> sets (restrict_space M (A \<inter> B))" | |
| 2536 | then obtain Y where "Y \<in> sets M" "X = Y \<inter> A \<inter> B" | |
| 2537 | by (auto simp: sets_restrict_space) | |
| 2538 | with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X" | |
| 2539 | by (subst (1 2) emeasure_restrict_space) | |
| 2540 | (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps) | |
| 2541 | qed | |
| 2542 | ||
| 2543 | lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \<inter> B)" | |
| 54417 | 2544 | proof (rule measure_eqI) | 
| 57025 | 2545 | show "sets (restrict_space (count_space B) A) = sets (count_space (A \<inter> B))" | 
| 2546 | by (subst sets_restrict_space) auto | |
| 54417 | 2547 | moreover fix X assume "X \<in> sets (restrict_space (count_space B) A)" | 
| 57025 | 2548 | ultimately have "X \<subseteq> A \<inter> B" by auto | 
| 2549 | then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \<inter> B)) X" | |
| 54417 | 2550 | by (cases "finite X") (auto simp add: emeasure_restrict_space) | 
| 2551 | qed | |
| 2552 | ||
| 60063 | 2553 | lemma sigma_finite_measure_restrict_space: | 
| 2554 | assumes "sigma_finite_measure M" | |
| 2555 | and A: "A \<in> sets M" | |
| 2556 | shows "sigma_finite_measure (restrict_space M A)" | |
| 2557 | proof - | |
| 2558 | interpret sigma_finite_measure M by fact | |
| 2559 | from sigma_finite_countable obtain C | |
| 2560 | where C: "countable C" "C \<subseteq> sets M" "(\<Union>C) = space M" "\<forall>a\<in>C. emeasure M a \<noteq> \<infinity>" | |
| 2561 | by blast | |
| 67399 | 2562 | let ?C = "(\<inter>) A ` C" | 
| 60063 | 2563 | from C have "countable ?C" "?C \<subseteq> sets (restrict_space M A)" "(\<Union>?C) = space (restrict_space M A)" | 
| 2564 | by(auto simp add: sets_restrict_space space_restrict_space) | |
| 2565 |   moreover {
 | |
| 2566 | fix a | |
| 2567 | assume "a \<in> ?C" | |
| 2568 | then obtain a' where "a = A \<inter> a'" "a' \<in> C" .. | |
| 2569 | then have "emeasure (restrict_space M A) a \<le> emeasure M a'" | |
| 2570 | using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono) | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2571 | also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by (simp add: less_top) | 
| 60063 | 2572 | finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp } | 
| 2573 | ultimately show ?thesis | |
| 2574 | by unfold_locales (rule exI conjI|assumption|blast)+ | |
| 2575 | qed | |
| 2576 | ||
| 2577 | lemma finite_measure_restrict_space: | |
| 2578 | assumes "finite_measure M" | |
| 2579 | and A: "A \<in> sets M" | |
| 2580 | shows "finite_measure (restrict_space M A)" | |
| 2581 | proof - | |
| 2582 | interpret finite_measure M by fact | |
| 2583 | show ?thesis | |
| 2584 | by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) | |
| 2585 | qed | |
| 2586 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2587 | lemma restrict_distr: | 
| 57137 | 2588 | assumes [measurable]: "f \<in> measurable M N" | 
| 2589 | assumes [simp]: "\<Omega> \<inter> space N \<in> sets N" and restrict: "f \<in> space M \<rightarrow> \<Omega>" | |
| 2590 | shows "restrict_space (distr M N f) \<Omega> = distr M (restrict_space N \<Omega>) f" | |
| 2591 | (is "?l = ?r") | |
| 2592 | proof (rule measure_eqI) | |
| 2593 | fix A assume "A \<in> sets (restrict_space (distr M N f) \<Omega>)" | |
| 2594 | with restrict show "emeasure ?l A = emeasure ?r A" | |
| 2595 | by (subst emeasure_distr) | |
| 2596 | (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr | |
| 2597 | intro!: measurable_restrict_space2) | |
| 2598 | qed (simp add: sets_restrict_space) | |
| 2599 | ||
| 59000 | 2600 | lemma measure_eqI_restrict_generator: | 
| 2601 | assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X" | |
| 2602 | assumes sets_eq: "sets M = sets N" and \<Omega>: "\<Omega> \<in> sets M" | |
| 2603 | assumes "sets (restrict_space M \<Omega>) = sigma_sets \<Omega> E" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2604 | assumes "sets (restrict_space N \<Omega>) = sigma_sets \<Omega> E" | 
| 59000 | 2605 | assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" | 
| 2606 |   assumes A: "countable A" "A \<noteq> {}" "A \<subseteq> E" "\<Union>A = \<Omega>" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
 | |
| 2607 | shows "M = N" | |
| 2608 | proof (rule measure_eqI) | |
| 2609 | fix X assume X: "X \<in> sets M" | |
| 2610 | then have "emeasure M X = emeasure (restrict_space M \<Omega>) (X \<inter> \<Omega>)" | |
| 2611 | using ae \<Omega> by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE) | |
| 2612 | also have "restrict_space M \<Omega> = restrict_space N \<Omega>" | |
| 2613 | proof (rule measure_eqI_generator_eq) | |
| 2614 | fix X assume "X \<in> E" | |
| 2615 | then show "emeasure (restrict_space M \<Omega>) X = emeasure (restrict_space N \<Omega>) X" | |
| 2616 | using E \<Omega> by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) | |
| 2617 | next | |
| 2618 | show "range (from_nat_into A) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>" | |
| 69546 
27dae626822b
prefer naming convention from datatype package for strong congruence rules
 haftmann parents: 
69541diff
changeset | 2619 | using A by (auto cong del: SUP_cong_simp) | 
| 59000 | 2620 | next | 
| 2621 | fix i | |
| 2622 | have "emeasure (restrict_space M \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)" | |
| 2623 | using A \<Omega> by (subst emeasure_restrict_space) | |
| 2624 | (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into) | |
| 2625 | with A show "emeasure (restrict_space M \<Omega>) (from_nat_into A i) \<noteq> \<infinity>" | |
| 2626 | by (auto intro: from_nat_into) | |
| 2627 | qed fact+ | |
| 2628 | also have "emeasure (restrict_space N \<Omega>) (X \<inter> \<Omega>) = emeasure N X" | |
| 2629 | using X ae \<Omega> by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE) | |
| 2630 | finally show "emeasure M X = emeasure N X" . | |
| 2631 | qed fact | |
| 2632 | ||
| 70136 | 2633 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Null measure\<close> | 
| 59425 | 2634 | |
| 69564 | 2635 | definition null_measure :: "'a measure \<Rightarrow> 'a measure" where | 
| 2636 | "null_measure M = sigma (space M) (sets M)" | |
| 59425 | 2637 | |
| 2638 | lemma space_null_measure[simp]: "space (null_measure M) = space M" | |
| 2639 | by (simp add: null_measure_def) | |
| 2640 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2641 | lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" | 
| 59425 | 2642 | by (simp add: null_measure_def) | 
| 2643 | ||
| 2644 | lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" | |
| 2645 | by (cases "X \<in> sets M", rule emeasure_measure_of) | |
| 2646 | (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def | |
| 2647 | dest: sets.sets_into_space) | |
| 2648 | ||
| 2649 | lemma measure_null_measure[simp]: "measure (null_measure M) X = 0" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2650 | by (intro measure_eq_emeasure_eq_ennreal) auto | 
| 59425 | 2651 | |
| 61633 | 2652 | lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2653 | by(rule measure_eqI) simp_all | 
| 61633 | 2654 | |
| 61634 | 2655 | subsection \<open>Scaling a measure\<close> | 
| 2656 | ||
| 70136 | 2657 | definition\<^marker>\<open>tag important\<close> scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where | 
| 69564 | 2658 | "scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)" | 
| 61634 | 2659 | |
| 2660 | lemma space_scale_measure: "space (scale_measure r M) = space M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2661 | by (simp add: scale_measure_def) | 
| 61634 | 2662 | |
| 2663 | lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2664 | by (simp add: scale_measure_def) | 
| 61634 | 2665 | |
| 2666 | lemma emeasure_scale_measure [simp]: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2667 | "emeasure (scale_measure r M) A = r * emeasure M A" | 
| 61634 | 2668 | (is "_ = ?\<mu> A") | 
| 2669 | proof(cases "A \<in> sets M") | |
| 2670 | case True | |
| 2671 | show ?thesis unfolding scale_measure_def | |
| 2672 | proof(rule emeasure_measure_of_sigma) | |
| 2673 | show "sigma_algebra (space M) (sets M)" .. | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2674 | show "positive (sets M) ?\<mu>" by (simp add: positive_def) | 
| 61634 | 2675 | show "countably_additive (sets M) ?\<mu>" | 
| 2676 | proof (rule countably_additiveI) | |
| 2677 | fix A :: "nat \<Rightarrow> _" assume *: "range A \<subseteq> sets M" "disjoint_family A" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2678 | have "(\<Sum>i. ?\<mu> (A i)) = r * (\<Sum>i. emeasure M (A i))" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2679 | by simp | 
| 61634 | 2680 | also have "\<dots> = ?\<mu> (\<Union>i. A i)" using * by(simp add: suminf_emeasure) | 
| 2681 | finally show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" . | |
| 2682 | qed | |
| 2683 | qed(fact True) | |
| 2684 | qed(simp add: emeasure_notin_sets) | |
| 2685 | ||
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2686 | lemma scale_measure_1 [simp]: "scale_measure 1 M = M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2687 | by(rule measure_eqI) simp_all | 
| 61634 | 2688 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2689 | lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2690 | by(rule measure_eqI) simp_all | 
| 61634 | 2691 | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2692 | lemma measure_scale_measure [simp]: "0 \<le> r \<Longrightarrow> measure (scale_measure r M) A = r * measure M A" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2693 | using emeasure_scale_measure[of r M A] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2694 | emeasure_eq_ennreal_measure[of M A] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2695 | measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A] | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2696 | by (cases "emeasure (scale_measure r M) A = top") | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2697 | (auto simp del: emeasure_scale_measure | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2698 | simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric]) | 
| 61634 | 2699 | |
| 2700 | lemma scale_scale_measure [simp]: | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2701 | "scale_measure r (scale_measure r' M) = scale_measure (r * r') M" | 
| 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2702 | by (rule measure_eqI) (simp_all add: max_def mult.assoc) | 
| 61634 | 2703 | |
| 2704 | lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M" | |
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 2705 | by (rule measure_eqI) simp_all | 
| 61634 | 2706 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2707 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2708 | subsection \<open>Complete lattice structure on measures\<close> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2709 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2710 | lemma (in finite_measure) finite_measure_Diff': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2711 | "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A - B) = measure M A - measure M (A \<inter> B)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2712 | using finite_measure_Diff[of A "A \<inter> B"] by (auto simp: Diff_Int) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2713 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2714 | lemma (in finite_measure) finite_measure_Union': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2715 | "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2716 | using finite_measure_Union[of A "B - A"] by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2717 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2718 | lemma finite_unsigned_Hahn_decomposition: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2719 | assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2720 |   shows "\<exists>Y\<in>sets M. (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2721 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2722 | interpret M: finite_measure M by fact | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2723 | interpret N: finite_measure N by fact | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2724 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2725 | define d where "d X = measure M X - measure N X" for X | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2726 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2727 | have [intro]: "bdd_above (d`sets M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2728 | using sets.sets_into_space[of _ M] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2729 | by (intro bdd_aboveI[where M="measure M (space M)"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2730 | (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2731 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2732 | define \<gamma> where "\<gamma> = (SUP X\<in>sets M. d X)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2733 | have le_\<gamma>[intro]: "X \<in> sets M \<Longrightarrow> d X \<le> \<gamma>" for X | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2734 | by (auto simp: \<gamma>_def intro!: cSUP_upper) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2735 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2736 | have "\<exists>f. \<forall>n. f n \<in> sets M \<and> d (f n) > \<gamma> - 1 / 2^n" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2737 | proof (intro choice_iff[THEN iffD1] allI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2738 | fix n | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2739 | have "\<exists>X\<in>sets M. \<gamma> - 1 / 2^n < d X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2740 | unfolding \<gamma>_def by (intro less_cSUP_iff[THEN iffD1]) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2741 | then show "\<exists>y. y \<in> sets M \<and> \<gamma> - 1 / 2 ^ n < d y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2742 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2743 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2744 | then obtain E where [measurable]: "E n \<in> sets M" and E: "d (E n) > \<gamma> - 1 / 2^n" for n | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2745 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2746 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2747 |   define F where "F m n = (if m \<le> n then \<Inter>i\<in>{m..n}. E i else space M)" for m n
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2748 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2749 | have [measurable]: "m \<le> n \<Longrightarrow> F m n \<in> sets M" for m n | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2750 | by (auto simp: F_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2751 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2752 | have 1: "\<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)" if "m \<le> n" for m n | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2753 | using that | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2754 | proof (induct rule: dec_induct) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2755 | case base with E[of m] show ?case | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2756 | by (simp add: F_def field_simps) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2757 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2758 | case (step i) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2759 | have F_Suc: "F m (Suc i) = F m i \<inter> E (Suc i)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2760 | using \<open>m \<le> i\<close> by (auto simp: F_def le_Suc_eq) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2761 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2762 | have "\<gamma> + (\<gamma> - 2 / 2^m + 1 / 2 ^ Suc i) \<le> (\<gamma> - 1 / 2^Suc i) + (\<gamma> - 2 / 2^m + 1 / 2^i)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2763 | by (simp add: field_simps) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2764 | also have "\<dots> \<le> d (E (Suc i)) + d (F m i)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2765 | using E[of "Suc i"] by (intro add_mono step) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2766 | also have "\<dots> = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2767 | using \<open>m \<le> i\<close> by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2768 | also have "\<dots> = d (E (Suc i) \<union> F m i) + d (F m (Suc i))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2769 | using \<open>m \<le> i\<close> by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2770 | also have "\<dots> \<le> \<gamma> + d (F m (Suc i))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2771 | using \<open>m \<le> i\<close> by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2772 | finally show ?case | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2773 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2774 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2775 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2776 |   define F' where "F' m = (\<Inter>i\<in>{m..}. E i)" for m
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2777 | have F'_eq: "F' m = (\<Inter>i. F m (i + m))" for m | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2778 | by (fastforce simp: le_iff_add[of m] F'_def F_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2779 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2780 | have [measurable]: "F' m \<in> sets M" for m | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2781 | by (auto simp: F'_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2782 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2783 | have \<gamma>_le: "\<gamma> - 0 \<le> d (\<Union>m. F' m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2784 | proof (rule LIMSEQ_le) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2785 | show "(\<lambda>n. \<gamma> - 2 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 0" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2786 | by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2787 | have "incseq F'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2788 | by (auto simp: incseq_def F'_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2789 | then show "(\<lambda>m. d (F' m)) \<longlonglongrightarrow> d (\<Union>m. F' m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2790 | unfolding d_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2791 | by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2792 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2793 | have "\<gamma> - 2 / 2 ^ m + 0 \<le> d (F' m)" for m | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2794 | proof (rule LIMSEQ_le) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2795 | have *: "decseq (\<lambda>n. F m (n + m))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2796 | by (auto simp: decseq_def F_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2797 | show "(\<lambda>n. d (F m n)) \<longlonglongrightarrow> d (F' m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2798 | unfolding d_def F'_eq | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2799 | by (rule LIMSEQ_offset[where k=m]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2800 | (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2801 | show "(\<lambda>n. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 2 / 2 ^ m + 0" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2802 | by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2803 | show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2804 | using 1[of m] by (intro exI[of _ m]) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2805 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2806 | then show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ n \<le> d (F' n)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2807 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2808 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2809 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2810 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2811 | proof (safe intro!: bexI[of _ "\<Union>m. F' m"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2812 | fix X assume [measurable]: "X \<in> sets M" and X: "X \<subseteq> (\<Union>m. F' m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2813 | have "d (\<Union>m. F' m) - d X = d ((\<Union>m. F' m) - X)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2814 | using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2815 | also have "\<dots> \<le> \<gamma>" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2816 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2817 | finally have "0 \<le> d X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2818 | using \<gamma>_le by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2819 | then show "emeasure N X \<le> emeasure M X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2820 | by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2821 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2822 |     fix X assume [measurable]: "X \<in> sets M" and X: "X \<inter> (\<Union>m. F' m) = {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2823 | then have "d (\<Union>m. F' m) + d X = d (X \<union> (\<Union>m. F' m))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2824 | by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2825 | also have "\<dots> \<le> \<gamma>" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2826 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2827 | finally have "d X \<le> 0" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2828 | using \<gamma>_le by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2829 | then show "emeasure M X \<le> emeasure N X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2830 | by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2831 | qed auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2832 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2833 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2834 | proposition unsigned_Hahn_decomposition: | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2835 | assumes [simp]: "sets N = sets M" and [measurable]: "A \<in> sets M" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2836 | and [simp]: "emeasure M A \<noteq> top" "emeasure N A \<noteq> top" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2837 |   shows "\<exists>Y\<in>sets M. Y \<subseteq> A \<and> (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<subseteq> A \<longrightarrow> X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
 | 
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2838 | proof - | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2839 | have "\<exists>Y\<in>sets (restrict_space M A). | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2840 | (\<forall>X\<in>sets (restrict_space M A). X \<subseteq> Y \<longrightarrow> (restrict_space N A) X \<le> (restrict_space M A) X) \<and> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2841 |     (\<forall>X\<in>sets (restrict_space M A). X \<inter> Y = {} \<longrightarrow> (restrict_space M A) X \<le> (restrict_space N A) X)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2842 | proof (rule finite_unsigned_Hahn_decomposition) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2843 | show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2844 | by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2845 | qed (simp add: sets_restrict_space) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2846 | then show ?thesis | 
| 74362 | 2847 | apply - | 
| 2848 | apply (erule bexE) | |
| 2849 | subgoal for Y | |
| 2850 | apply (intro bexI[of _ Y] conjI ballI conjI) | |
| 2851 | apply (simp_all add: sets_restrict_space emeasure_restrict_space) | |
| 2852 | apply safe | |
| 2853 | subgoal for X Z | |
| 2854 | by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1) | |
| 2855 | subgoal for X Z | |
| 2856 | by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps) | |
| 2857 | apply auto | |
| 2858 | done | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2859 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2860 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2861 | |
| 70136 | 2862 | text\<^marker>\<open>tag important\<close> \<open> | 
| 69597 | 2863 | Define a lexicographical order on \<^type>\<open>measure\<close>, in the order space, sets and measure. The parts | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2864 | of the lexicographical order are point-wise ordered. | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2865 | \<close> | 
| 60772 | 2866 | |
| 68617 | 2867 | instantiation measure :: (type) order_bot | 
| 60772 | 2868 | begin | 
| 2869 | ||
| 2870 | inductive less_eq_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2871 | "space M \<subset> space N \<Longrightarrow> less_eq_measure M N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2872 | | "space M = space N \<Longrightarrow> sets M \<subset> sets N \<Longrightarrow> less_eq_measure M N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2873 | | "space M = space N \<Longrightarrow> sets M = sets N \<Longrightarrow> emeasure M \<le> emeasure N \<Longrightarrow> less_eq_measure M N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2874 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2875 | lemma le_measure_iff: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2876 | "M \<le> N \<longleftrightarrow> (if space M = space N then | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2877 | if sets M = sets N then emeasure M \<le> emeasure N else sets M \<subseteq> sets N else space M \<subseteq> space N)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2878 | by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) | 
| 60772 | 2879 | |
| 70136 | 2880 | definition\<^marker>\<open>tag important\<close> less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where | 
| 60772 | 2881 | "less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)" | 
| 2882 | ||
| 70136 | 2883 | definition\<^marker>\<open>tag important\<close> bot_measure :: "'a measure" where | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2884 |   "bot_measure = sigma {} {}"
 | 
| 60772 | 2885 | |
| 2886 | lemma | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2887 |   shows space_bot[simp]: "space bot = {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2888 |     and sets_bot[simp]: "sets bot = {{}}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2889 | and emeasure_bot[simp]: "emeasure bot X = 0" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2890 | by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2891 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2892 | instance | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2893 | proof standard | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2894 | show "bot \<le> a" for a :: "'a measure" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2895 | by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2896 | qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2897 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2898 | end | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2899 | |
| 68607 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2900 | proposition le_measure: "sets M = sets N \<Longrightarrow> M \<le> N \<longleftrightarrow> (\<forall>A\<in>sets M. emeasure M A \<le> emeasure N A)" | 
| 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 immler parents: 
68532diff
changeset | 2901 | apply - | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2902 | apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2903 | subgoal for X | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2904 | by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2905 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2906 | |
| 70136 | 2907 | definition\<^marker>\<open>tag important\<close> sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where | 
| 69564 | 2908 | "sup_measure' A B = | 
| 2909 | measure_of (space A) (sets A) | |
| 2910 | (\<lambda>X. SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))" | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2911 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2912 | lemma assumes [simp]: "sets B = sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2913 | shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2914 | and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2915 | using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2916 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2917 | lemma emeasure_sup_measure': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2918 | assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \<in> sets A" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2919 | shows "emeasure (sup_measure' A B) X = (SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2920 | (is "_ = ?S X") | 
| 60772 | 2921 | proof - | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2922 | note sets_eq_imp_space_eq[OF sets_eq, simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2923 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2924 | using sup_measure'_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2925 | proof (rule emeasure_measure_of) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2926 | let ?d = "\<lambda>X Y. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2927 | show "countably_additive (sets (sup_measure' A B)) (\<lambda>X. SUP Y \<in> sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2928 | proof (rule countably_additiveI, goal_cases) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2929 | case (1 X) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2930 | then have [measurable]: "\<And>i. X i \<in> sets A" and "disjoint_family X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2931 | by auto | 
| 69661 | 2932 | have disjoint: "disjoint_family (\<lambda>i. X i \<inter> Y)" "disjoint_family (\<lambda>i. X i - Y)" for Y | 
| 2933 | by (auto intro: disjoint_family_on_bisimulation [OF \<open>disjoint_family X\<close>, simplified]) | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 2934 | have "(\<Sum>i. ?S (X i)) = (SUP Y\<in>sets A. \<Sum>i. ?d (X i) Y)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2935 | proof (rule ennreal_suminf_SUP_eq_directed) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2936 | fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \<in> sets A" "b \<in> sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2937 | have "\<exists>c\<in>sets A. c \<subseteq> X i \<and> (\<forall>a\<in>sets A. ?d (X i) a \<le> ?d (X i) c)" for i | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2938 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2939 | assume "emeasure A (X i) = top \<or> emeasure B (X i) = top" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2940 | then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2941 | proof | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2942 | assume "emeasure A (X i) = top" then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2943 | by (intro bexI[of _ "X i"]) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2944 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2945 | assume "emeasure B (X i) = top" then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2946 |               by (intro bexI[of _ "{}"]) auto
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2947 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2948 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2949 | assume finite: "\<not> (emeasure A (X i) = top \<or> emeasure B (X i) = top)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2950 |           then have "\<exists>Y\<in>sets A. Y \<subseteq> X i \<and> (\<forall>C\<in>sets A. C \<subseteq> Y \<longrightarrow> B C \<le> A C) \<and> (\<forall>C\<in>sets A. C \<subseteq> X i \<longrightarrow> C \<inter> Y = {} \<longrightarrow> A C \<le> B C)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2951 | using unsigned_Hahn_decomposition[of B A "X i"] by simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2952 | then obtain Y where [measurable]: "Y \<in> sets A" and [simp]: "Y \<subseteq> X i" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2953 | and B_le_A: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> B C \<le> A C" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2954 |             and A_le_B: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> X i \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> A C \<le> B C"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2955 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2956 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2957 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2958 | proof (intro bexI[of _ Y] ballI conjI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2959 | fix a assume [measurable]: "a \<in> sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2960 | have *: "(X i \<inter> a \<inter> Y \<union> (X i \<inter> a - Y)) = X i \<inter> a" "(X i - a) \<inter> Y \<union> (X i - a - Y) = X i \<inter> - a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2961 | for a Y by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2962 | then have "?d (X i) a = | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2963 | (A (X i \<inter> a \<inter> Y) + A (X i \<inter> a \<inter> - Y)) + (B (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2964 | by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2965 | also have "\<dots> \<le> (A (X i \<inter> a \<inter> Y) + B (X i \<inter> a \<inter> - Y)) + (A (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2966 | by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2967 | also have "\<dots> \<le> (A (X i \<inter> Y \<inter> a) + A (X i \<inter> Y \<inter> - a)) + (B (X i \<inter> - Y \<inter> a) + B (X i \<inter> - Y \<inter> - a))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2968 | by (simp add: ac_simps) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2969 | also have "\<dots> \<le> A (X i \<inter> Y) + B (X i \<inter> - Y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2970 | by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2971 | finally show "?d (X i) a \<le> ?d (X i) Y" . | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2972 | qed auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2973 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2974 | then obtain C where [measurable]: "C i \<in> sets A" and "C i \<subseteq> X i" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2975 | and C: "\<And>a. a \<in> sets A \<Longrightarrow> ?d (X i) a \<le> ?d (X i) (C i)" for i | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2976 | by metis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2977 | have *: "X i \<inter> (\<Union>i. C i) = X i \<inter> C i" for i | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2978 | proof safe | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2979 | fix x j assume "x \<in> X i" "x \<in> C j" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2980 |           moreover have "i = j \<or> X i \<inter> X j = {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2981 | using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2982 | ultimately show "x \<in> C i" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2983 | using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2984 | qed auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2985 | have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2986 | proof safe | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2987 | fix x j assume "x \<in> X i" "x \<notin> C i" "x \<in> C j" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2988 |           moreover have "i = j \<or> X i \<inter> X j = {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2989 | using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2990 | ultimately show False | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2991 | using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2992 | qed auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2993 | show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2994 | apply (intro bexI[of _ "\<Union>i. C i"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2995 | unfolding * ** | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2996 | apply (intro C ballI conjI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2997 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2998 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 2999 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3000 | also have "\<dots> = ?S (\<Union>i. X i)" | 
| 69661 | 3001 | apply (simp only: UN_extend_simps(4)) | 
| 3002 | apply (rule arg_cong [of _ _ Sup]) | |
| 3003 | apply (rule image_cong) | |
| 3004 | apply (fact refl) | |
| 3005 | using disjoint | |
| 3006 | apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps) | |
| 3007 | done | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3008 | finally show "(\<Sum>i. ?S (X i)) = ?S (\<Union>i. X i)" . | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3009 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3010 | qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const) | 
| 60772 | 3011 | qed | 
| 3012 | ||
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3013 | lemma le_emeasure_sup_measure'1: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3014 | assumes "sets B = sets A" "X \<in> sets A" shows "emeasure A X \<le> emeasure (sup_measure' A B) X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3015 | by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3016 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3017 | lemma le_emeasure_sup_measure'2: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3018 | assumes "sets B = sets A" "X \<in> sets A" shows "emeasure B X \<le> emeasure (sup_measure' A B) X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3019 |   by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3020 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3021 | lemma emeasure_sup_measure'_le2: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3022 | assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \<in> sets C" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3023 | assumes A: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure A Y \<le> emeasure C Y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3024 | assumes B: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure B Y \<le> emeasure C Y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3025 | shows "emeasure (sup_measure' A B) X \<le> emeasure C X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3026 | proof (subst emeasure_sup_measure') | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3027 | show "(SUP Y\<in>sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)) \<le> emeasure C X" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3028 | unfolding \<open>sets A = sets C\<close> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3029 | proof (intro SUP_least) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3030 | fix Y assume [measurable]: "Y \<in> sets C" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3031 | have [simp]: "X \<inter> Y \<union> (X - Y) = X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3032 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3033 | have "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C (X \<inter> Y) + emeasure C (X \<inter> - Y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3034 | by (intro add_mono A B) (auto simp: Diff_eq[symmetric]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3035 | also have "\<dots> = emeasure C X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3036 | by (subst plus_emeasure) (auto simp: Diff_eq[symmetric]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3037 | finally show "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C X" . | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3038 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3039 | qed simp_all | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3040 | |
| 70136 | 3041 | definition\<^marker>\<open>tag important\<close> sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 69564 | 3042 | "sup_lexord A B k s c = | 
| 3043 | (if k A = k B then c else | |
| 3044 | if \<not> k A \<le> k B \<and> \<not> k B \<le> k A then s else | |
| 3045 | if k B \<le> k A then A else B)" | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3046 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3047 | lemma sup_lexord: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3048 | "(k A < k B \<Longrightarrow> P B) \<Longrightarrow> (k B < k A \<Longrightarrow> P A) \<Longrightarrow> (k A = k B \<Longrightarrow> P c) \<Longrightarrow> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3049 | (\<not> k B \<le> k A \<Longrightarrow> \<not> k A \<le> k B \<Longrightarrow> P s) \<Longrightarrow> P (sup_lexord A B k s c)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3050 | by (auto simp: sup_lexord_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3051 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3052 | lemmas le_sup_lexord = sup_lexord[where P="\<lambda>a. c \<le> a" for c] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3053 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3054 | lemma sup_lexord1: "k A = k B \<Longrightarrow> sup_lexord A B k s c = c" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3055 | by (simp add: sup_lexord_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3056 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3057 | lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3058 | by (auto simp: sup_lexord_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3059 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3060 | lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \<A> \<subseteq> sets x) = (\<A> \<subseteq> sets x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3061 | using sets.sigma_sets_subset[of \<A> x] by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3062 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3063 | lemma sigma_le_iff: "\<A> \<subseteq> Pow \<Omega> \<Longrightarrow> sigma \<Omega> \<A> \<le> x \<longleftrightarrow> (\<Omega> \<subseteq> space x \<and> (space x = \<Omega> \<longrightarrow> \<A> \<subseteq> sets x))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3064 | by (cases "\<Omega> = space x") | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3065 | (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3066 | sigma_sets_superset_generator sigma_sets_le_sets_iff) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3067 | |
| 68617 | 3068 | instantiation measure :: (type) semilattice_sup | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3069 | begin | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3070 | |
| 70136 | 3071 | definition\<^marker>\<open>tag important\<close> sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3072 | "sup_measure A B = | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3073 |     sup_lexord A B space (sigma (space A \<union> space B) {})
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3074 | (sup_lexord A B sets (sigma (space A) (sets A \<union> sets B)) (sup_measure' A B))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3075 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3076 | instance | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3077 | proof | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3078 | fix x y z :: "'a measure" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3079 | show "x \<le> sup x y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3080 | unfolding sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3081 | proof (intro le_sup_lexord) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3082 | assume "space x = space y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3083 | then have *: "sets x \<union> sets y \<subseteq> Pow (space x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3084 | using sets.space_closed by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3085 | assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3086 | then have "sets x \<subset> sets x \<union> sets y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3087 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3088 | also have "\<dots> \<le> sigma (space x) (sets x \<union> sets y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3089 | by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3090 | finally show "x \<le> sigma (space x) (sets x \<union> sets y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3091 | by (simp add: space_measure_of[OF *] less_eq_measure.intros(2)) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3092 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3093 | assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3094 |     then show "x \<le> sigma (space x \<union> space y) {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3095 | by (intro less_eq_measure.intros) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3096 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3097 | assume "sets x = sets y" then show "x \<le> sup_measure' x y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3098 | by (simp add: le_measure le_emeasure_sup_measure'1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3099 | qed (auto intro: less_eq_measure.intros) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3100 | show "y \<le> sup x y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3101 | unfolding sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3102 | proof (intro le_sup_lexord) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3103 | assume **: "space x = space y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3104 | then have *: "sets x \<union> sets y \<subseteq> Pow (space y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3105 | using sets.space_closed by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3106 | assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3107 | then have "sets y \<subset> sets x \<union> sets y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3108 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3109 | also have "\<dots> \<le> sigma (space y) (sets x \<union> sets y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3110 | by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3111 | finally show "y \<le> sigma (space x) (sets x \<union> sets y)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3112 | by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2)) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3113 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3114 | assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3115 |     then show "y \<le> sigma (space x \<union> space y) {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3116 | by (intro less_eq_measure.intros) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3117 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3118 | assume "sets x = sets y" then show "y \<le> sup_measure' x y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3119 | by (simp add: le_measure le_emeasure_sup_measure'2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3120 | qed (auto intro: less_eq_measure.intros) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3121 | show "x \<le> y \<Longrightarrow> z \<le> y \<Longrightarrow> sup x z \<le> y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3122 | unfolding sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3123 | proof (intro sup_lexord[where P="\<lambda>x. x \<le> y"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3124 | assume "x \<le> y" "z \<le> y" and [simp]: "space x = space z" "sets x = sets z" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3125 | from \<open>x \<le> y\<close> show "sup_measure' x z \<le> y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3126 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3127 | case 1 then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3128 | by (intro less_eq_measure.intros(1)) simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3129 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3130 | case 2 then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3131 | by (intro less_eq_measure.intros(2)) simp_all | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3132 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3133 | case 3 with \<open>z \<le> y\<close> \<open>x \<le> y\<close> show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3134 | by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3135 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3136 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3137 | assume **: "x \<le> y" "z \<le> y" "space x = space z" "\<not> sets z \<subseteq> sets x" "\<not> sets x \<subseteq> sets z" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3138 | then have *: "sets x \<union> sets z \<subseteq> Pow (space x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3139 | using sets.space_closed by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3140 | show "sigma (space x) (sets x \<union> sets z) \<le> y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3141 | unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3142 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3143 | assume "x \<le> y" "z \<le> y" "\<not> space z \<subseteq> space x" "\<not> space x \<subseteq> space z" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3144 | then have "space x \<subseteq> space y" "space z \<subseteq> space y" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3145 | by (auto simp: le_measure_iff split: if_split_asm) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3146 |     then show "sigma (space x \<union> space z) {} \<le> y"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3147 | by (simp add: sigma_le_iff) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3148 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3149 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3150 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3151 | end | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3152 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3153 | lemma space_empty_eq_bot: "space a = {} \<longleftrightarrow> a = bot"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3154 | using space_empty[of a] by (auto intro!: measure_eqI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3155 | |
| 63657 | 3156 | lemma sets_eq_iff_bounded: "A \<le> B \<Longrightarrow> B \<le> C \<Longrightarrow> sets A = sets C \<Longrightarrow> sets B = sets A" | 
| 3157 | by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm) | |
| 3158 | ||
| 3159 | lemma sets_sup: "sets A = sets M \<Longrightarrow> sets B = sets M \<Longrightarrow> sets (sup A B) = sets M" | |
| 3160 | by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq) | |
| 3161 | ||
| 3162 | lemma le_measureD1: "A \<le> B \<Longrightarrow> space A \<le> space B" | |
| 3163 | by (auto simp: le_measure_iff split: if_split_asm) | |
| 3164 | ||
| 3165 | lemma le_measureD2: "A \<le> B \<Longrightarrow> space A = space B \<Longrightarrow> sets A \<le> sets B" | |
| 3166 | by (auto simp: le_measure_iff split: if_split_asm) | |
| 3167 | ||
| 3168 | lemma le_measureD3: "A \<le> B \<Longrightarrow> sets A = sets B \<Longrightarrow> emeasure A X \<le> emeasure B X" | |
| 3169 | by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm) | |
| 3170 | ||
| 69313 | 3171 | lemma UN_space_closed: "\<Union>(sets ` S) \<subseteq> Pow (\<Union>(space ` S))" | 
| 63657 | 3172 | using sets.space_closed by auto | 
| 3173 | ||
| 70136 | 3174 | definition\<^marker>\<open>tag important\<close> | 
| 69564 | 3175 |   Sup_lexord :: "('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> 'a"
 | 
| 63657 | 3176 | where | 
| 69564 | 3177 | "Sup_lexord k c s A = | 
| 3178 | (let U = (SUP a\<in>A. k a) | |
| 3179 |    in if \<exists>a\<in>A. k a = U then c {a\<in>A. k a = U} else s A)"
 | |
| 63657 | 3180 | |
| 3181 | lemma Sup_lexord: | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3182 |   "(\<And>a S. a \<in> A \<Longrightarrow> k a = (SUP a\<in>A. k a) \<Longrightarrow> S = {a'\<in>A. k a' = k a} \<Longrightarrow> P (c S)) \<Longrightarrow> ((\<And>a. a \<in> A \<Longrightarrow> k a \<noteq> (SUP a\<in>A. k a)) \<Longrightarrow> P (s A)) \<Longrightarrow>
 | 
| 63657 | 3183 | P (Sup_lexord k c s A)" | 
| 3184 | by (auto simp: Sup_lexord_def Let_def) | |
| 3185 | ||
| 3186 | lemma Sup_lexord1: | |
| 3187 |   assumes A: "A \<noteq> {}" "(\<And>a. a \<in> A \<Longrightarrow> k a = (\<Union>a\<in>A. k a))" "P (c A)"
 | |
| 3188 | shows "P (Sup_lexord k c s A)" | |
| 3189 | unfolding Sup_lexord_def Let_def | |
| 3190 | proof (clarsimp, safe) | |
| 3191 | show "\<forall>a\<in>A. k a \<noteq> (\<Union>x\<in>A. k x) \<Longrightarrow> P (s A)" | |
| 3192 | by (metis assms(1,2) ex_in_conv) | |
| 3193 | next | |
| 3194 | fix a assume "a \<in> A" "k a = (\<Union>x\<in>A. k x)" | |
| 3195 |   then have "{a \<in> A. k a = (\<Union>x\<in>A. k x)} = {a \<in> A. k a = k a}"
 | |
| 3196 | by (metis A(2)[symmetric]) | |
| 3197 |   then show "P (c {a \<in> A. k a = (\<Union>x\<in>A. k x)})"
 | |
| 3198 | by (simp add: A(3)) | |
| 3199 | qed | |
| 3200 | ||
| 68617 | 3201 | instantiation measure :: (type) complete_lattice | 
| 63658 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3202 | begin | 
| 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3203 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3204 | interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" | 
| 63658 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3205 | by standard (auto intro!: antisym) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3206 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3207 | lemma sup_measure_F_mono': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3208 | "finite J \<Longrightarrow> finite I \<Longrightarrow> sup_measure.F id I \<le> sup_measure.F id (I \<union> J)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3209 | proof (induction J rule: finite_induct) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3210 | case empty then show ?case | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3211 | by simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3212 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3213 | case (insert i J) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3214 | show ?case | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3215 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3216 | assume "i \<in> I" with insert show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3217 | by (auto simp: insert_absorb) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3218 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3219 | assume "i \<notin> I" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3220 | have "sup_measure.F id I \<le> sup_measure.F id (I \<union> J)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3221 | by (intro insert) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3222 | also have "\<dots> \<le> sup_measure.F id (insert i (I \<union> J))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3223 | using insert \<open>i \<notin> I\<close> by (subst sup_measure.insert) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3224 | finally show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3225 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3226 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3227 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3228 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3229 | lemma sup_measure_F_mono: "finite I \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sup_measure.F id J \<le> sup_measure.F id I" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3230 | using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3231 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3232 | lemma sets_sup_measure_F: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3233 |   "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> sets i = sets M) \<Longrightarrow> sets (sup_measure.F id I) = sets M"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3234 | by (induction I rule: finite_ne_induct) (simp_all add: sets_sup) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3235 | |
| 70136 | 3236 | definition\<^marker>\<open>tag important\<close> Sup_measure' :: "'a measure set \<Rightarrow> 'a measure" where | 
| 69564 | 3237 | "Sup_measure' M = | 
| 3238 | measure_of (\<Union>a\<in>M. space a) (\<Union>a\<in>M. sets a) | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3239 |     (\<lambda>X. (SUP P\<in>{P. finite P \<and> P \<subseteq> M }. sup_measure.F id P X))"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3240 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3241 | lemma space_Sup_measure'2: "space (Sup_measure' M) = (\<Union>m\<in>M. space m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3242 | unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3243 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3244 | lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3245 | unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3246 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3247 | lemma sets_Sup_measure': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3248 |   assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3249 | shows "sets (Sup_measure' M) = sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3250 |   using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close> by (simp add: Sup_measure'_def)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3251 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3252 | lemma space_Sup_measure': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3253 |   assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3254 | shows "space (Sup_measure' M) = space A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3255 |   using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close>
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3256 | by (simp add: Sup_measure'_def ) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3257 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3258 | lemma emeasure_Sup_measure': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3259 |   assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "X \<in> sets A" "M \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3260 |   shows "emeasure (Sup_measure' M) X = (SUP P\<in>{P. finite P \<and> P \<subseteq> M}. sup_measure.F id P X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3261 | (is "_ = ?S X") | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3262 | using Sup_measure'_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3263 | proof (rule emeasure_measure_of) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3264 | note sets_eq[THEN sets_eq_imp_space_eq, simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3265 | have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3266 |     using \<open>M \<noteq> {}\<close> by (simp_all add: Sup_measure'_def)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3267 | let ?\<mu> = "sup_measure.F id" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3268 | show "countably_additive (sets (Sup_measure' M)) ?S" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3269 | proof (rule countably_additiveI, goal_cases) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3270 | case (1 F) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3271 | then have **: "range F \<subseteq> sets A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3272 | by (auto simp: *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3273 | show "(\<Sum>i. ?S (F i)) = ?S (\<Union>i. F i)" | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 3274 | proof (subst ennreal_suminf_SUP_eq_directed) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3275 |       fix i j and N :: "nat set" assume ij: "i \<in> {P. finite P \<and> P \<subseteq> M}" "j \<in> {P. finite P \<and> P \<subseteq> M}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3276 |       have "(i \<noteq> {} \<longrightarrow> sets (?\<mu> i) = sets A) \<and> (j \<noteq> {} \<longrightarrow> sets (?\<mu> j) = sets A) \<and>
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3277 |         (i \<noteq> {} \<or> j \<noteq> {} \<longrightarrow> sets (?\<mu> (i \<union> j)) = sets A)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3278 | using ij by (intro impI sets_sup_measure_F conjI) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3279 | then have "?\<mu> j (F n) \<le> ?\<mu> (i \<union> j) (F n) \<and> ?\<mu> i (F n) \<le> ?\<mu> (i \<union> j) (F n)" for n | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3280 | using ij | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3281 |         by (cases "i = {}"; cases "j = {}")
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3282 | (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3283 | simp del: id_apply) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3284 |       with ij show "\<exists>k\<in>{P. finite P \<and> P \<subseteq> M}. \<forall>n\<in>N. ?\<mu> i (F n) \<le> ?\<mu> k (F n) \<and> ?\<mu> j (F n) \<le> ?\<mu> k (F n)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3285 | by (safe intro!: bexI[of _ "i \<union> j"]) auto | 
| 60772 | 3286 | next | 
| 69313 | 3287 |       show "(SUP P \<in> {P. finite P \<and> P \<subseteq> M}. \<Sum>n. ?\<mu> P (F n)) = (SUP P \<in> {P. finite P \<and> P \<subseteq> M}. ?\<mu> P (\<Union>(F ` UNIV)))"
 | 
| 69661 | 3288 | proof (intro arg_cong [of _ _ Sup] image_cong refl) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3289 |         fix i assume i: "i \<in> {P. finite P \<and> P \<subseteq> M}"
 | 
| 69313 | 3290 | show "(\<Sum>n. ?\<mu> i (F n)) = ?\<mu> i (\<Union>(F ` UNIV))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3291 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3292 |           assume "i \<noteq> {}" with i ** show ?thesis
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3293 | apply (intro suminf_emeasure \<open>disjoint_family F\<close>) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3294 | apply (subst sets_sup_measure_F[OF _ _ sets_eq]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3295 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3296 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3297 | qed simp | 
| 60772 | 3298 | qed | 
| 62975 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 hoelzl parents: 
62390diff
changeset | 3299 | qed | 
| 60772 | 3300 | qed | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3301 | show "positive (sets (Sup_measure' M)) ?S" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3302 | by (auto simp: positive_def bot_ennreal[symmetric]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3303 | show "X \<in> sets (Sup_measure' M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3304 | using assms * by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3305 | qed (rule UN_space_closed) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3306 | |
| 70136 | 3307 | definition\<^marker>\<open>tag important\<close> Sup_measure :: "'a measure set \<Rightarrow> 'a measure" where | 
| 69564 | 3308 | "Sup_measure = | 
| 3309 | Sup_lexord space | |
| 3310 | (Sup_lexord sets Sup_measure' | |
| 3311 | (\<lambda>U. sigma (\<Union>u\<in>U. space u) (\<Union>u\<in>U. sets u))) | |
| 3312 |     (\<lambda>U. sigma (\<Union>u\<in>U. space u) {})"
 | |
| 3313 | ||
| 70136 | 3314 | definition\<^marker>\<open>tag important\<close> Inf_measure :: "'a measure set \<Rightarrow> 'a measure" where | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3315 |   "Inf_measure A = Sup {x. \<forall>a\<in>A. x \<le> a}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3316 | |
| 70136 | 3317 | definition\<^marker>\<open>tag important\<close> inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" where | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3318 |   "inf_measure a b = Inf {a, b}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3319 | |
| 70136 | 3320 | definition\<^marker>\<open>tag important\<close> top_measure :: "'a measure" where | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3321 |   "top_measure = Inf {}"
 | 
| 60772 | 3322 | |
| 3323 | instance | |
| 3324 | proof | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3325 | note UN_space_closed [simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3326 | show upper: "x \<le> Sup A" if x: "x \<in> A" for x :: "'a measure" and A | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3327 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3328 | proof (intro Sup_lexord[where P="\<lambda>y. x \<le> y"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3329 | assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3330 |     from this[OF \<open>x \<in> A\<close>] \<open>x \<in> A\<close> show "x \<le> sigma (\<Union>a\<in>A. space a) {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3331 | by (intro less_eq_measure.intros) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3332 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3333 |     fix a S assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3334 | and neq: "\<And>aa. aa \<in> S \<Longrightarrow> sets aa \<noteq> (\<Union>a\<in>S. sets a)" | 
| 69313 | 3335 | have sp_a: "space a = (\<Union>(space ` S))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3336 | using \<open>a\<in>A\<close> by (auto simp: S) | 
| 69313 | 3337 | show "x \<le> sigma (\<Union>(space ` S)) (\<Union>(sets ` S))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3338 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3339 | assume [simp]: "space x = space a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3340 | have "sets x \<subset> (\<Union>a\<in>S. sets a)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3341 | using \<open>x\<in>A\<close> neq[of x] by (auto simp: S) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3342 | also have "\<dots> \<subseteq> sigma_sets (\<Union>x\<in>S. space x) (\<Union>x\<in>S. sets x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3343 | by (rule sigma_sets_superset_generator) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3344 | finally show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3345 | by (intro less_eq_measure.intros(2)) (simp_all add: sp_a) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3346 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3347 | assume "space x \<noteq> space a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3348 | moreover have "space x \<le> space a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3349 | unfolding a using \<open>x\<in>A\<close> by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3350 | ultimately show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3351 | by (intro less_eq_measure.intros) (simp add: less_le sp_a) | 
| 60772 | 3352 | qed | 
| 3353 | next | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3354 |     fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3355 |       and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3356 |     then have "S' \<noteq> {}" "space b = space a"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3357 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3358 | have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3359 | by (auto simp: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3360 | note sets_eq[THEN sets_eq_imp_space_eq, simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3361 | have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3362 |       using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3363 | show "x \<le> Sup_measure' S'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3364 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3365 | assume "x \<in> S" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3366 | with \<open>b \<in> S\<close> have "space x = space b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3367 | by (simp add: S) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3368 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3369 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3370 | assume "x \<in> S'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3371 | show "x \<le> Sup_measure' S'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3372 | proof (intro le_measure[THEN iffD2] ballI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3373 | show "sets x = sets (Sup_measure' S')" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3374 | using \<open>x\<in>S'\<close> * by (simp add: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3375 | fix X assume "X \<in> sets x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3376 | show "emeasure x X \<le> emeasure (Sup_measure' S') X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3377 | proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets x\<close>]) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3378 |             show "emeasure x X \<le> (SUP P \<in> {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3379 |               using \<open>x\<in>S'\<close> by (intro SUP_upper2[where i="{x}"]) auto
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3380 | qed (insert \<open>x\<in>S'\<close> S', auto) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3381 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3382 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3383 | assume "x \<notin> S'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3384 | then have "sets x \<noteq> sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3385 | using \<open>x\<in>S\<close> by (auto simp: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3386 | moreover have "sets x \<le> sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3387 | using \<open>x\<in>S\<close> unfolding b by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3388 | ultimately show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3389 | using * \<open>x \<in> S\<close> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3390 | by (intro less_eq_measure.intros(2)) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3391 | (simp_all add: * \<open>space x = space b\<close> less_le) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3392 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3393 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3394 | assume "x \<notin> S" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3395 | with \<open>x\<in>A\<close> \<open>x \<notin> S\<close> \<open>space b = space a\<close> show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3396 | by (intro less_eq_measure.intros) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3397 | (simp_all add: * less_le a SUP_upper S) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3398 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3399 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3400 | show least: "Sup A \<le> x" if x: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x" for x :: "'a measure" and A | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3401 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3402 | proof (intro Sup_lexord[where P="\<lambda>y. y \<le> x"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3403 | assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)" | 
| 69313 | 3404 |     show "sigma (\<Union>(space ` A)) {} \<le> x"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3405 | using x[THEN le_measureD1] by (subst sigma_le_iff) auto | 
| 60772 | 3406 | next | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3407 |     fix a S assume "a \<in> A" "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3408 | "\<And>a. a \<in> S \<Longrightarrow> sets a \<noteq> (\<Union>a\<in>S. sets a)" | 
| 69313 | 3409 | have "\<Union>(space ` S) \<subseteq> space x" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3410 | using S le_measureD1[OF x] by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3411 | moreover | 
| 69313 | 3412 | have "\<Union>(space ` S) = space a" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3413 | using \<open>a\<in>A\<close> S by auto | 
| 69313 | 3414 | then have "space x = \<Union>(space ` S) \<Longrightarrow> \<Union>(sets ` S) \<subseteq> sets x" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3415 | using \<open>a \<in> A\<close> le_measureD2[OF x] by (auto simp: S) | 
| 69313 | 3416 | ultimately show "sigma (\<Union>(space ` S)) (\<Union>(sets ` S)) \<le> x" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3417 | by (subst sigma_le_iff) simp_all | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3418 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3419 |     fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3420 |       and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3421 |     then have "S' \<noteq> {}" "space b = space a"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3422 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3423 | have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3424 | by (auto simp: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3425 | note sets_eq[THEN sets_eq_imp_space_eq, simp] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3426 | have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3427 |       using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3428 | show "Sup_measure' S' \<le> x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3429 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3430 | assume "space x = space a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3431 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3432 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3433 | assume **: "sets x = sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3434 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3435 | proof (intro le_measure[THEN iffD2] ballI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3436 | show ***: "sets (Sup_measure' S') = sets x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3437 | by (simp add: * **) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3438 | fix X assume "X \<in> sets (Sup_measure' S')" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3439 | show "emeasure (Sup_measure' S') X \<le> emeasure x X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3440 | unfolding *** | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3441 | proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets (Sup_measure' S')\<close>]) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3442 |             show "(SUP P \<in> {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X) \<le> emeasure x X"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3443 | proof (safe intro!: SUP_least) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3444 | fix P assume P: "finite P" "P \<subseteq> S'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3445 | show "emeasure (sup_measure.F id P) X \<le> emeasure x X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3446 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3447 |                 assume "P = {}" then show ?thesis
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3448 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3449 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3450 |                 assume "P \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3451 | from P have "finite P" "P \<subseteq> A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3452 | unfolding S' S by (simp_all add: subset_eq) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3453 | then have "sup_measure.F id P \<le> x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3454 | by (induction P) (auto simp: x) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3455 | moreover have "sets (sup_measure.F id P) = sets x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3456 |                   using \<open>finite P\<close> \<open>P \<noteq> {}\<close> \<open>P \<subseteq> S'\<close> \<open>sets x = sets b\<close>
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3457 | by (intro sets_sup_measure_F) (auto simp: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3458 | ultimately show "emeasure (sup_measure.F id P) X \<le> emeasure x X" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3459 | by (rule le_measureD3) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3460 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3461 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3462 | show "m \<in> S' \<Longrightarrow> sets m = sets (Sup_measure' S')" for m | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3463 | unfolding * by (simp add: S') | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3464 | qed fact | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3465 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3466 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3467 | assume "sets x \<noteq> sets b" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3468 | moreover have "sets b \<le> sets x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3469 | unfolding b S using x[THEN le_measureD2] \<open>space x = space a\<close> by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3470 | ultimately show "Sup_measure' S' \<le> x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3471 | using \<open>space x = space a\<close> \<open>b \<in> S\<close> | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3472 | by (intro less_eq_measure.intros(2)) (simp_all add: * S) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3473 | qed | 
| 60772 | 3474 | next | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3475 | assume "space x \<noteq> space a" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3476 | then have "space a < space x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3477 | using le_measureD1[OF x[OF \<open>a\<in>A\<close>]] by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3478 | then show "Sup_measure' S' \<le> x" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3479 | by (intro less_eq_measure.intros) (simp add: * \<open>space b = space a\<close>) | 
| 60772 | 3480 | qed | 
| 3481 | qed | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3482 |   show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3483 | by (auto intro!: antisym least simp: top_measure_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3484 | show lower: "x \<in> A \<Longrightarrow> Inf A \<le> x" for x :: "'a measure" and A | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3485 | unfolding Inf_measure_def by (intro least) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3486 | show greatest: "(\<And>z. z \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> x \<le> Inf A" for x :: "'a measure" and A | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3487 | unfolding Inf_measure_def by (intro upper) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3488 | show "inf x y \<le> x" "inf x y \<le> y" "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> inf y z" for x y z :: "'a measure" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3489 | by (auto simp: inf_measure_def intro!: lower greatest) | 
| 60772 | 3490 | qed | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3491 | |
| 60772 | 3492 | end | 
| 3493 | ||
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3494 | lemma sets_SUP: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3495 | assumes "\<And>x. x \<in> I \<Longrightarrow> sets (M x) = sets N" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3496 |   shows "I \<noteq> {} \<Longrightarrow> sets (SUP i\<in>I. M i) = sets N"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3497 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3498 | using assms assms[THEN sets_eq_imp_space_eq] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3499 | sets_Sup_measure'[where A=N and M="M`I"] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3500 | by (intro Sup_lexord1[where P="\<lambda>x. sets x = sets N"]) auto | 
| 61633 | 3501 | |
| 3502 | lemma emeasure_SUP: | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3503 |   assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N" "I \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3504 |   shows "emeasure (SUP i\<in>I. M i) X = (SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emeasure (SUP i\<in>J. M i) X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3505 | proof - | 
| 63658 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3506 | interpret sup_measure: comm_monoid_set sup "bot :: 'b measure" | 
| 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3507 | by standard (auto intro!: antisym) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3508 | have eq: "finite J \<Longrightarrow> sup_measure.F id J = (SUP i\<in>J. i)" for J :: "'b measure set" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3509 | by (induction J rule: finite_induct) auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3510 |   have 1: "J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sets (SUP x\<in>J. M x) = sets N" for J
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3511 | by (intro sets_SUP sets) (auto ) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3512 |   from \<open>I \<noteq> {}\<close> obtain i where "i\<in>I" by auto
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3513 |   have "Sup_measure' (M`I) X = (SUP P\<in>{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3514 | using sets by (intro emeasure_Sup_measure') auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3515 | also have "Sup_measure' (M`I) = (SUP i\<in>I. M i)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3516 |     unfolding Sup_measure_def using \<open>I \<noteq> {}\<close> sets sets(1)[THEN sets_eq_imp_space_eq]
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3517 | by (intro Sup_lexord1[where P="\<lambda>x. _ = x"]) auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3518 |   also have "(SUP P\<in>{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X) =
 | 
| 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3519 |     (SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. (SUP i\<in>J. M i) X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3520 | proof (intro SUP_eq) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3521 |     fix J assume "J \<in> {P. finite P \<and> P \<subseteq> M`I}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3522 | then obtain J' where J': "J' \<subseteq> I" "finite J'" and J: "J = M`J'" and "finite J" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3523 | using finite_subset_image[of J M I] by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3524 |     show "\<exists>j\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. sup_measure.F id J X \<le> (SUP i\<in>j. M i) X"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3525 | proof cases | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3526 |       assume "J' = {}" with \<open>i \<in> I\<close> show ?thesis
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3527 | by (auto simp add: J) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3528 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3529 |       assume "J' \<noteq> {}" with J J' show ?thesis
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3530 | by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3531 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3532 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3533 |     fix J assume J: "J \<in> {P. P \<noteq> {} \<and> finite P \<and> P \<subseteq> I}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3534 |     show "\<exists>J'\<in>{J. finite J \<and> J \<subseteq> M`I}. (SUP i\<in>J. M i) X \<le> sup_measure.F id J' X"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3535 | using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3536 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3537 | finally show ?thesis . | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3538 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3539 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3540 | lemma emeasure_SUP_chain: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3541 | assumes sets: "\<And>i. i \<in> A \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N" | 
| 67399 | 3542 |   assumes ch: "Complete_Partial_Order.chain (\<le>) (M ` A)" and "A \<noteq> {}"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3543 | shows "emeasure (SUP i\<in>A. M i) X = (SUP i\<in>A. emeasure (M i) X)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3544 | proof (subst emeasure_SUP[OF sets \<open>A \<noteq> {}\<close>])
 | 
| 69313 | 3545 |   show "(SUP J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (Sup (M ` J)) X) = (SUP i\<in>A. emeasure (M i) X)"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3546 | proof (rule SUP_eq) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3547 |     fix J assume "J \<in> {J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}"
 | 
| 67399 | 3548 |     then have J: "Complete_Partial_Order.chain (\<le>) (M ` J)" "finite J" "J \<noteq> {}" and "J \<subseteq> A"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3549 | using ch[THEN chain_subset, of "M`J"] by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3550 | with in_chain_finite[OF J(1)] obtain j where "j \<in> J" "(SUP j\<in>J. M j) = M j" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3551 | by auto | 
| 69313 | 3552 | with \<open>J \<subseteq> A\<close> show "\<exists>j\<in>A. emeasure (Sup (M ` J)) X \<le> emeasure (M j) X" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3553 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3554 | next | 
| 69313 | 3555 |     fix j assume "j\<in>A" then show "\<exists>i\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (M j) X \<le> emeasure (Sup (M ` i)) X"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3556 |       by (intro bexI[of _ "{j}"]) auto
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3557 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3558 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3559 | |
| 70136 | 3560 | subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Supremum of a set of \<open>\<sigma>\<close>-algebras\<close> | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3561 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3562 | lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3563 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3564 | apply (intro Sup_lexord[where P="\<lambda>x. space x = _"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3565 | apply (subst space_Sup_measure'2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3566 | apply auto [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3567 | apply (subst space_measure_of[OF UN_space_closed]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3568 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3569 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3570 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3571 | lemma sets_Sup_eq: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3572 |   assumes *: "\<And>m. m \<in> M \<Longrightarrow> space m = X" and "M \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3573 | shows "sets (Sup M) = sigma_sets X (\<Union>x\<in>M. sets x)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3574 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3575 | apply (rule Sup_lexord1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3576 | apply fact | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3577 | apply (simp add: assms) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3578 | apply (rule Sup_lexord) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3579 | subgoal premises that for a S | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3580 | unfolding that(3) that(2)[symmetric] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3581 | using that(1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3582 | apply (subst sets_Sup_measure'2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3583 | apply (intro arg_cong2[where f=sigma_sets]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3584 | apply (auto simp: *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3585 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3586 | apply (subst sets_measure_of[OF UN_space_closed]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3587 | apply (simp add: assms) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3588 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3589 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3590 | lemma in_sets_Sup: "(\<And>m. m \<in> M \<Longrightarrow> space m = X) \<Longrightarrow> m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3591 | by (subst sets_Sup_eq[where X=X]) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3592 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3593 | lemma Sup_lexord_rel: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3594 | assumes "\<And>i. i \<in> I \<Longrightarrow> k (A i) = k (B i)" | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3595 |     "R (c (A ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))})) (c (B ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))}))"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3596 | "R (s (A`I)) (s (B`I))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3597 | shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3598 | proof - | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3599 |   have "A ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))} =  {a \<in> A ` I. k a = (SUP x\<in>I. k (B x))}"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3600 | using assms(1) by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3601 |   moreover have "B ` {a \<in> I. k (B a) = (SUP x\<in>I. k (B x))} =  {a \<in> B ` I. k a = (SUP x\<in>I. k (B x))}"
 | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3602 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3603 | ultimately show ?thesis | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 3604 | using assms by (auto simp: Sup_lexord_def Let_def image_comp) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3605 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3606 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3607 | lemma sets_SUP_cong: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3608 | assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (SUP i\<in>I. M i) = sets (SUP i\<in>I. N i)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3609 | unfolding Sup_measure_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3610 | using eq eq[THEN sets_eq_imp_space_eq] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3611 | apply (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3612 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3613 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3614 | apply (simp add: sets_Sup_measure'2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3615 | apply (intro arg_cong2[where f="\<lambda>x y. sets (sigma x y)"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3616 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3617 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3618 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3619 | lemma sets_Sup_in_sets: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3620 |   assumes "M \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3621 | assumes "\<And>m. m \<in> M \<Longrightarrow> space m = space N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3622 | assumes "\<And>m. m \<in> M \<Longrightarrow> sets m \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3623 | shows "sets (Sup M) \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3624 | proof - | 
| 69313 | 3625 | have *: "\<Union>(space ` M) = space N" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3626 | using assms by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3627 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3628 | unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3629 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3630 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3631 | lemma measurable_Sup1: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3632 | assumes m: "m \<in> M" and f: "f \<in> measurable m N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3633 | and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3634 | shows "f \<in> measurable (Sup M) N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3635 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3636 | have "space (Sup M) = space m" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3637 | using m by (auto simp add: space_Sup_eq_UN dest: const_space) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3638 | then show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3639 | using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3640 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3641 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3642 | lemma measurable_Sup2: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3643 |   assumes M: "M \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3644 | assumes f: "\<And>m. m \<in> M \<Longrightarrow> f \<in> measurable N m" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3645 | and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3646 | shows "f \<in> measurable N (Sup M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3647 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3648 | from M obtain m where "m \<in> M" by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3649 | have space_eq: "\<And>n. n \<in> M \<Longrightarrow> space n = space m" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3650 | by (intro const_space \<open>m \<in> M\<close>) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3651 | have "f \<in> measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m))" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3652 | proof (rule measurable_measure_of) | 
| 69313 | 3653 | show "f \<in> space N \<rightarrow> \<Union>(space ` M)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3654 | using measurable_space[OF f] M by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3655 | qed (auto intro: measurable_sets f dest: sets.sets_into_space) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3656 | also have "measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)) = measurable N (Sup M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3657 | apply (intro measurable_cong_sets refl) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3658 | apply (subst sets_Sup_eq[OF space_eq M]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3659 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3660 | apply (subst sets_measure_of[OF UN_space_closed]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3661 | apply (simp add: space_eq M) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3662 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3663 | finally show ?thesis . | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3664 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3665 | |
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 3666 | lemma measurable_SUP2: | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 3667 |   "I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f \<in> measurable N (M i)) \<Longrightarrow>
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3668 | (\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> space (M i) = space (M j)) \<Longrightarrow> f \<in> measurable N (SUP i\<in>I. M i)" | 
| 64320 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 3669 | by (auto intro!: measurable_Sup2) | 
| 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 hoelzl parents: 
64283diff
changeset | 3670 | |
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3671 | lemma sets_Sup_sigma: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3672 |   assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3673 | shows "sets (SUP m\<in>M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3674 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3675 |   { fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3676 | then have "a \<in> sigma_sets \<Omega> (\<Union>M)" | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 3677 | by induction (auto intro: sigma_sets.intros(2-)) } | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3678 | then show "sets (SUP m\<in>M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3679 | apply (subst sets_Sup_eq[where X="\<Omega>"]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3680 | apply (auto simp add: M) [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3681 | apply auto [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3682 | apply (simp add: space_measure_of_conv M Union_least) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3683 | apply (rule sigma_sets_eqI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3684 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3685 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3686 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3687 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3688 | lemma Sup_sigma: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3689 |   assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3690 | shows "(SUP m\<in>M. sigma \<Omega> m) = (sigma \<Omega> (\<Union>M))" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3691 | proof (intro antisym SUP_least) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3692 | have *: "\<Union>M \<subseteq> Pow \<Omega>" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3693 | using M by auto | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3694 | show "sigma \<Omega> (\<Union>M) \<le> (SUP m\<in>M. sigma \<Omega> m)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3695 | proof (intro less_eq_measure.intros(3)) | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3696 | show "space (sigma \<Omega> (\<Union>M)) = space (SUP m\<in>M. sigma \<Omega> m)" | 
| 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3697 | "sets (sigma \<Omega> (\<Union>M)) = sets (SUP m\<in>M. sigma \<Omega> m)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3698 | using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3699 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3700 | qed (simp add: emeasure_sigma le_fun_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3701 | fix m assume "m \<in> M" then show "sigma \<Omega> m \<le> sigma \<Omega> (\<Union>M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3702 | by (subst sigma_le_iff) (auto simp add: M *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3703 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3704 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3705 | lemma SUP_sigma_sigma: | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3706 |   "M \<noteq> {} \<Longrightarrow> (\<And>m. m \<in> M \<Longrightarrow> f m \<subseteq> Pow \<Omega>) \<Longrightarrow> (SUP m\<in>M. sigma \<Omega> (f m)) = sigma \<Omega> (\<Union>m\<in>M. f m)"
 | 
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 3707 | using Sup_sigma[of "f`M" \<Omega>] by (auto simp: image_comp) | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3708 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3709 | lemma sets_vimage_Sup_eq: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3710 |   assumes *: "M \<noteq> {}" "f \<in> X \<rightarrow> Y" "\<And>m. m \<in> M \<Longrightarrow> space m = Y"
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3711 | shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \<in> M. vimage_algebra X f m)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3712 | (is "?IS = ?SI") | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3713 | proof | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3714 | show "?IS \<subseteq> ?SI" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3715 | apply (intro sets_image_in_sets measurable_Sup2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3716 | apply (simp add: space_Sup_eq_UN *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3717 | apply (simp add: *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3718 | apply (intro measurable_Sup1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3719 | apply (rule imageI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3720 | apply assumption | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3721 | apply (rule measurable_vimage_algebra1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3722 | apply (auto simp: *) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3723 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3724 | show "?SI \<subseteq> ?IS" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3725 | apply (intro sets_Sup_in_sets) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3726 | apply (auto simp: *) [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3727 | apply (auto simp: *) [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3728 | apply (elim imageE) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3729 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3730 | apply (rule sets_image_in_sets) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3731 | apply simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3732 | apply (simp add: measurable_def) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3733 | apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3734 | apply (auto intro: in_sets_Sup[OF *(3)]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3735 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3736 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3737 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3738 | lemma restrict_space_eq_vimage_algebra': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3739 | "sets (restrict_space M \<Omega>) = sets (vimage_algebra (\<Omega> \<inter> space M) (\<lambda>x. x) M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3740 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3741 |   have *: "{A \<inter> (\<Omega> \<inter> space M) |A. A \<in> sets M} = {A \<inter> \<Omega> |A. A \<in> sets M}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3742 | using sets.sets_into_space[of _ M] by blast | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3743 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3744 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3745 | unfolding restrict_space_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3746 | by (subst sets_measure_of) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3747 | (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3748 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3749 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3750 | lemma sigma_le_sets: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3751 | assumes [simp]: "A \<subseteq> Pow X" shows "sets (sigma X A) \<subseteq> sets N \<longleftrightarrow> X \<in> sets N \<and> A \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3752 | proof | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3753 | have "X \<in> sigma_sets X A" "A \<subseteq> sigma_sets X A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3754 | by (auto intro: sigma_sets_top) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3755 | moreover assume "sets (sigma X A) \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3756 | ultimately show "X \<in> sets N \<and> A \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3757 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3758 | next | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3759 | assume *: "X \<in> sets N \<and> A \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3760 |   { fix Y assume "Y \<in> sigma_sets X A" from this * have "Y \<in> sets N"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3761 | by induction auto } | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3762 | then show "sets (sigma X A) \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3763 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3764 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3765 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3766 | lemma measurable_iff_sets: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3767 | "f \<in> measurable M N \<longleftrightarrow> (f \<in> space M \<rightarrow> space N \<and> sets (vimage_algebra (space M) f N) \<subseteq> sets M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3768 | proof - | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3769 |   have *: "{f -` A \<inter> space M |A. A \<in> sets N} \<subseteq> Pow (space M)"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3770 | by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3771 | show ?thesis | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3772 | unfolding measurable_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3773 | by (auto simp add: vimage_algebra_def sigma_le_sets[OF *]) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3774 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3775 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3776 | lemma sets_vimage_algebra_space: "X \<in> sets (vimage_algebra X f M)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3777 | using sets.top[of "vimage_algebra X f M"] by simp | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3778 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3779 | lemma measurable_mono: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3780 | assumes N: "sets N' \<le> sets N" "space N = space N'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3781 | assumes M: "sets M \<le> sets M'" "space M = space M'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3782 | shows "measurable M N \<subseteq> measurable M' N'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3783 | unfolding measurable_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3784 | proof safe | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3785 | fix f A assume "f \<in> space M \<rightarrow> space N" "A \<in> sets N'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3786 | moreover assume "\<forall>y\<in>sets N. f -` y \<inter> space M \<in> sets M" note this[THEN bspec, of A] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3787 | ultimately show "f -` A \<inter> space M' \<in> sets M'" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3788 | using assms by auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3789 | qed (insert N M, auto) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3790 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3791 | lemma measurable_Sup_measurable: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3792 | assumes f: "f \<in> space N \<rightarrow> A" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3793 |   shows "f \<in> measurable N (Sup {M. space M = A \<and> f \<in> measurable N M})"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3794 | proof (rule measurable_Sup2) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3795 |   show "{M. space M = A \<and> f \<in> measurable N M} \<noteq> {}"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3796 | using f unfolding ex_in_conv[symmetric] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3797 |     by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of)
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3798 | qed auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3799 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3800 | lemma (in sigma_algebra) sigma_sets_subset': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3801 | assumes a: "a \<subseteq> M" "\<Omega>' \<in> M" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3802 | shows "sigma_sets \<Omega>' a \<subseteq> M" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3803 | proof | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3804 | show "x \<in> M" if x: "x \<in> sigma_sets \<Omega>' a" for x | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3805 | using x by (induct rule: sigma_sets.induct) (insert a, auto) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3806 | qed | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3807 | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3808 | lemma in_sets_SUP: "i \<in> I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> space (M i) = Y) \<Longrightarrow> X \<in> sets (M i) \<Longrightarrow> X \<in> sets (SUP i\<in>I. M i)" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3809 | by (intro in_sets_Sup[where X=Y]) auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3810 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3811 | lemma measurable_SUP1: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3812 | "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<And>m n. m \<in> I \<Longrightarrow> n \<in> I \<Longrightarrow> space (M m) = space (M n)) \<Longrightarrow> | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69164diff
changeset | 3813 | f \<in> measurable (SUP i\<in>I. M i) N" | 
| 63333 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3814 | by (auto intro: measurable_Sup1) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3815 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3816 | lemma sets_image_in_sets': | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3817 | assumes X: "X \<in> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3818 | assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3819 | shows "sets (vimage_algebra X f M) \<subseteq> sets N" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3820 | unfolding sets_vimage_algebra | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3821 | by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3822 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3823 | lemma mono_vimage_algebra: | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3824 | "sets M \<le> sets N \<Longrightarrow> sets (vimage_algebra X f M) \<subseteq> sets (vimage_algebra X f N)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3825 |   using sets.top[of "sigma X {f -` A \<inter> X |A. A \<in> sets N}"]
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3826 | unfolding vimage_algebra_def | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3827 | apply (subst (asm) space_measure_of) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3828 | apply auto [] | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3829 | apply (subst sigma_le_sets) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3830 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3831 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3832 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3833 | lemma mono_restrict_space: "sets M \<le> sets N \<Longrightarrow> sets (restrict_space M X) \<subseteq> sets (restrict_space N X)" | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3834 | unfolding sets_restrict_space by (rule image_mono) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3835 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3836 | lemma sets_eq_bot: "sets M = {{}} \<longleftrightarrow> M = bot"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3837 | apply safe | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3838 | apply (intro measure_eqI) | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3839 | apply auto | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3840 | done | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3841 | |
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3842 | lemma sets_eq_bot2: "{{}} = sets M \<longleftrightarrow> M = bot"
 | 
| 
158ab2239496
Probability: show that measures form a complete lattice
 hoelzl parents: 
63040diff
changeset | 3843 | using sets_eq_bot[of M] by blast | 
| 61633 | 3844 | |
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3845 | |
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3846 | lemma (in finite_measure) countable_support: | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3847 |   "countable {x. measure M {x} \<noteq> 0}"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3848 | proof cases | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3849 | assume "measure M (space M) = 0" | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3850 |   with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3851 | by auto | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3852 | then show ?thesis | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3853 | by simp | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3854 | next | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3855 |   let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3856 | assume "?M \<noteq> 0" | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3857 |   then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3858 | using reals_Archimedean[of "?m x / ?M" for x] | 
| 63958 
02de4a58e210
HOL-Analysis: add measurable sets with finite measures, prove affine transformation rule for the Lebesgue measure
 hoelzl parents: 
63940diff
changeset | 3859 | by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff) | 
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3860 |   have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3861 | proof (rule ccontr) | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3862 |     fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3863 | then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X" | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3864 | by (metis infinite_arbitrarily_large) | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3865 | from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x" | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3866 | by auto | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3867 |     { fix x assume "x \<in> X"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3868 | from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff) | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3869 |       then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3870 | note singleton_sets = this | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3871 | have "?M < (\<Sum>x\<in>X. ?M / Suc n)" | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3872 | using \<open>?M \<noteq> 0\<close> | 
| 63658 
7faa9bf9860b
epheremal interpretation keeps auxiliary definition localized
 haftmann parents: 
63657diff
changeset | 3873 | by (simp add: \<open>card X = Suc (Suc n)\<close> field_simps less_le) | 
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3874 | also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)" | 
| 64267 | 3875 | by (rule sum_mono) fact | 
| 63626 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3876 |     also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3877 | using singleton_sets \<open>finite X\<close> | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3878 | by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3879 |     finally have "?M < measure M (\<Union>x\<in>X. {x})" .
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3880 |     moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
 | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3881 | using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3882 | ultimately show False by simp | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3883 | qed | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3884 | show ?thesis | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3885 | unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3886 | qed | 
| 
44ce6b524ff3
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
 hoelzl parents: 
63540diff
changeset | 3887 | |
| 60772 | 3888 | end |