| 47694 |      1 | (*  Title:      HOL/Probability/Measure_Space.thy
 | 
|  |      2 |     Author:     Lawrence C Paulson
 | 
|  |      3 |     Author:     Johannes Hölzl, TU München
 | 
|  |      4 |     Author:     Armin Heller, TU München
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | header {* Measure spaces and their properties *}
 | 
|  |      8 | 
 | 
|  |      9 | theory Measure_Space
 | 
|  |     10 | imports
 | 
|  |     11 |   Sigma_Algebra
 | 
|  |     12 |   "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
 | 
|  |     13 | begin
 | 
|  |     14 | 
 | 
|  |     15 | lemma suminf_cmult_indicator:
 | 
|  |     16 |   fixes f :: "nat \<Rightarrow> ereal"
 | 
|  |     17 |   assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
 | 
|  |     18 |   shows "(\<Sum>n. f n * indicator (A n) x) = f i"
 | 
|  |     19 | proof -
 | 
|  |     20 |   have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
 | 
|  |     21 |     using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
 | 
|  |     22 |   then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
 | 
|  |     23 |     by (auto simp: setsum_cases)
 | 
|  |     24 |   moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
 | 
|  |     25 |   proof (rule ereal_SUPI)
 | 
|  |     26 |     fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
 | 
|  |     27 |     from this[of "Suc i"] show "f i \<le> y" by auto
 | 
|  |     28 |   qed (insert assms, simp)
 | 
|  |     29 |   ultimately show ?thesis using assms
 | 
|  |     30 |     by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
 | 
|  |     31 | qed
 | 
|  |     32 | 
 | 
|  |     33 | lemma suminf_indicator:
 | 
|  |     34 |   assumes "disjoint_family A"
 | 
|  |     35 |   shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
 | 
|  |     36 | proof cases
 | 
|  |     37 |   assume *: "x \<in> (\<Union>i. A i)"
 | 
|  |     38 |   then obtain i where "x \<in> A i" by auto
 | 
|  |     39 |   from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
 | 
|  |     40 |   show ?thesis using * by simp
 | 
|  |     41 | qed simp
 | 
|  |     42 | 
 | 
|  |     43 | text {*
 | 
|  |     44 |   The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
 | 
|  |     45 |   represent sigma algebras (with an arbitrary emeasure).
 | 
|  |     46 | *}
 | 
|  |     47 | 
 | 
|  |     48 | section "Extend binary sets"
 | 
|  |     49 | 
 | 
|  |     50 | lemma LIMSEQ_binaryset:
 | 
|  |     51 |   assumes f: "f {} = 0"
 | 
|  |     52 |   shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
 | 
|  |     53 | proof -
 | 
|  |     54 |   have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
 | 
|  |     55 |     proof
 | 
|  |     56 |       fix n
 | 
|  |     57 |       show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
 | 
|  |     58 |         by (induct n)  (auto simp add: binaryset_def f)
 | 
|  |     59 |     qed
 | 
|  |     60 |   moreover
 | 
|  |     61 |   have "... ----> f A + f B" by (rule tendsto_const)
 | 
|  |     62 |   ultimately
 | 
|  |     63 |   have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
 | 
|  |     64 |     by metis
 | 
|  |     65 |   hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
 | 
|  |     66 |     by simp
 | 
|  |     67 |   thus ?thesis by (rule LIMSEQ_offset [where k=2])
 | 
|  |     68 | qed
 | 
|  |     69 | 
 | 
|  |     70 | lemma binaryset_sums:
 | 
|  |     71 |   assumes f: "f {} = 0"
 | 
|  |     72 |   shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
 | 
|  |     73 |     by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
 | 
|  |     74 | 
 | 
|  |     75 | lemma suminf_binaryset_eq:
 | 
|  |     76 |   fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
 | 
|  |     77 |   shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
 | 
|  |     78 |   by (metis binaryset_sums sums_unique)
 | 
|  |     79 | 
 | 
|  |     80 | section {* Properties of a premeasure @{term \<mu>} *}
 | 
|  |     81 | 
 | 
|  |     82 | text {*
 | 
|  |     83 |   The definitions for @{const positive} and @{const countably_additive} should be here, by they are
 | 
|  |     84 |   necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
 | 
|  |     85 | *}
 | 
|  |     86 | 
 | 
|  |     87 | definition additive where
 | 
|  |     88 |   "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
 | 
|  |     89 | 
 | 
|  |     90 | definition increasing where
 | 
|  |     91 |   "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
 | 
|  |     92 | 
 | 
|  |     93 | lemma positiveD_empty:
 | 
|  |     94 |   "positive M f \<Longrightarrow> f {} = 0"
 | 
|  |     95 |   by (auto simp add: positive_def)
 | 
|  |     96 | 
 | 
|  |     97 | lemma additiveD:
 | 
|  |     98 |   "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
 | 
|  |     99 |   by (auto simp add: additive_def)
 | 
|  |    100 | 
 | 
|  |    101 | lemma increasingD:
 | 
|  |    102 |   "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
 | 
|  |    103 |   by (auto simp add: increasing_def)
 | 
|  |    104 | 
 | 
|  |    105 | lemma countably_additiveI:
 | 
|  |    106 |   "(\<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))
 | 
|  |    107 |   \<Longrightarrow> countably_additive M f"
 | 
|  |    108 |   by (simp add: countably_additive_def)
 | 
|  |    109 | 
 | 
|  |    110 | lemma (in ring_of_sets) disjointed_additive:
 | 
|  |    111 |   assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
 | 
|  |    112 |   shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
 | 
|  |    113 | proof (induct n)
 | 
|  |    114 |   case (Suc n)
 | 
|  |    115 |   then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
 | 
|  |    116 |     by simp
 | 
|  |    117 |   also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
 | 
|  |    118 |     using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
 | 
|  |    119 |   also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
 | 
|  |    120 |     using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
 | 
|  |    121 |   finally show ?case .
 | 
|  |    122 | qed simp
 | 
|  |    123 | 
 | 
|  |    124 | lemma (in ring_of_sets) additive_sum:
 | 
|  |    125 |   fixes A:: "'i \<Rightarrow> 'a set"
 | 
|  |    126 |   assumes f: "positive M f" and ad: "additive M f" and "finite S"
 | 
|  |    127 |       and A: "A`S \<subseteq> M"
 | 
|  |    128 |       and disj: "disjoint_family_on A S"
 | 
|  |    129 |   shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
 | 
|  |    130 | using `finite S` disj A proof induct
 | 
|  |    131 |   case empty show ?case using f by (simp add: positive_def)
 | 
|  |    132 | next
 | 
|  |    133 |   case (insert s S)
 | 
|  |    134 |   then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
 | 
|  |    135 |     by (auto simp add: disjoint_family_on_def neq_iff)
 | 
|  |    136 |   moreover
 | 
|  |    137 |   have "A s \<in> M" using insert by blast
 | 
|  |    138 |   moreover have "(\<Union>i\<in>S. A i) \<in> M"
 | 
|  |    139 |     using insert `finite S` by auto
 | 
|  |    140 |   moreover
 | 
|  |    141 |   ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
 | 
|  |    142 |     using ad UNION_in_sets A by (auto simp add: additive_def)
 | 
|  |    143 |   with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
 | 
|  |    144 |     by (auto simp add: additive_def subset_insertI)
 | 
|  |    145 | qed
 | 
|  |    146 | 
 | 
|  |    147 | lemma (in ring_of_sets) additive_increasing:
 | 
|  |    148 |   assumes posf: "positive M f" and addf: "additive M f"
 | 
|  |    149 |   shows "increasing M f"
 | 
|  |    150 | proof (auto simp add: increasing_def)
 | 
|  |    151 |   fix x y
 | 
|  |    152 |   assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
 | 
|  |    153 |   then have "y - x \<in> M" by auto
 | 
|  |    154 |   then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
 | 
|  |    155 |   then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
 | 
|  |    156 |   also have "... = f (x \<union> (y-x))" using addf
 | 
|  |    157 |     by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
 | 
|  |    158 |   also have "... = f y"
 | 
|  |    159 |     by (metis Un_Diff_cancel Un_absorb1 xy(3))
 | 
|  |    160 |   finally show "f x \<le> f y" by simp
 | 
|  |    161 | qed
 | 
|  |    162 | 
 | 
|  |    163 | lemma (in ring_of_sets) countably_additive_additive:
 | 
|  |    164 |   assumes posf: "positive M f" and ca: "countably_additive M f"
 | 
|  |    165 |   shows "additive M f"
 | 
|  |    166 | proof (auto simp add: additive_def)
 | 
|  |    167 |   fix x y
 | 
|  |    168 |   assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
 | 
|  |    169 |   hence "disjoint_family (binaryset x y)"
 | 
|  |    170 |     by (auto simp add: disjoint_family_on_def binaryset_def)
 | 
|  |    171 |   hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
 | 
|  |    172 |          (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
 | 
|  |    173 |          f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
 | 
|  |    174 |     using ca
 | 
|  |    175 |     by (simp add: countably_additive_def)
 | 
|  |    176 |   hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
 | 
|  |    177 |          f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
 | 
|  |    178 |     by (simp add: range_binaryset_eq UN_binaryset_eq)
 | 
|  |    179 |   thus "f (x \<union> y) = f x + f y" using posf x y
 | 
|  |    180 |     by (auto simp add: Un suminf_binaryset_eq positive_def)
 | 
|  |    181 | qed
 | 
|  |    182 | 
 | 
|  |    183 | lemma (in algebra) increasing_additive_bound:
 | 
|  |    184 |   fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
 | 
|  |    185 |   assumes f: "positive M f" and ad: "additive M f"
 | 
|  |    186 |       and inc: "increasing M f"
 | 
|  |    187 |       and A: "range A \<subseteq> M"
 | 
|  |    188 |       and disj: "disjoint_family A"
 | 
|  |    189 |   shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
 | 
|  |    190 | proof (safe intro!: suminf_bound)
 | 
|  |    191 |   fix N
 | 
|  |    192 |   note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
 | 
|  |    193 |   have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
 | 
|  |    194 |     using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
 | 
|  |    195 |   also have "... \<le> f \<Omega>" using space_closed A
 | 
|  |    196 |     by (intro increasingD[OF inc] finite_UN) auto
 | 
|  |    197 |   finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
 | 
|  |    198 | qed (insert f A, auto simp: positive_def)
 | 
|  |    199 | 
 | 
|  |    200 | lemma (in ring_of_sets) countably_additiveI_finite:
 | 
|  |    201 |   assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
 | 
|  |    202 |   shows "countably_additive M \<mu>"
 | 
|  |    203 | proof (rule countably_additiveI)
 | 
|  |    204 |   fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
 | 
|  |    205 | 
 | 
|  |    206 |   have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
 | 
|  |    207 |   from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
 | 
|  |    208 | 
 | 
|  |    209 |   have inj_f: "inj_on f {i. F i \<noteq> {}}"
 | 
|  |    210 |   proof (rule inj_onI, simp)
 | 
|  |    211 |     fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
 | 
|  |    212 |     then have "f i \<in> F i" "f j \<in> F j" using f by force+
 | 
|  |    213 |     with disj * show "i = j" by (auto simp: disjoint_family_on_def)
 | 
|  |    214 |   qed
 | 
|  |    215 |   have "finite (\<Union>i. F i)"
 | 
|  |    216 |     by (metis F(2) assms(1) infinite_super sets_into_space)
 | 
|  |    217 | 
 | 
|  |    218 |   have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
 | 
|  |    219 |     by (auto simp: positiveD_empty[OF `positive M \<mu>`])
 | 
|  |    220 |   moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
 | 
|  |    221 |   proof (rule finite_imageD)
 | 
|  |    222 |     from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
 | 
|  |    223 |     then show "finite (f`{i. F i \<noteq> {}})"
 | 
|  |    224 |       by (rule finite_subset) fact
 | 
|  |    225 |   qed fact
 | 
|  |    226 |   ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
 | 
|  |    227 |     by (rule finite_subset)
 | 
|  |    228 | 
 | 
|  |    229 |   have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
 | 
|  |    230 |     using disj by (auto simp: disjoint_family_on_def)
 | 
|  |    231 | 
 | 
|  |    232 |   from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
 | 
| 47761 |    233 |     by (rule suminf_finite) auto
 | 
| 47694 |    234 |   also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
 | 
|  |    235 |     using fin_not_empty F_subset by (rule setsum_mono_zero_left) auto
 | 
|  |    236 |   also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
 | 
|  |    237 |     using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
 | 
|  |    238 |   also have "\<dots> = \<mu> (\<Union>i. F i)"
 | 
|  |    239 |     by (rule arg_cong[where f=\<mu>]) auto
 | 
|  |    240 |   finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
 | 
|  |    241 | qed
 | 
|  |    242 | 
 | 
|  |    243 | section {* Properties of @{const emeasure} *}
 | 
|  |    244 | 
 | 
|  |    245 | lemma emeasure_positive: "positive (sets M) (emeasure M)"
 | 
|  |    246 |   by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
 | 
|  |    247 | 
 | 
|  |    248 | lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
 | 
|  |    249 |   using emeasure_positive[of M] by (simp add: positive_def)
 | 
|  |    250 | 
 | 
|  |    251 | lemma emeasure_nonneg[intro!]: "0 \<le> emeasure M A"
 | 
|  |    252 |   using emeasure_notin_sets[of A M] emeasure_positive[of M]
 | 
|  |    253 |   by (cases "A \<in> sets M") (auto simp: positive_def)
 | 
|  |    254 | 
 | 
|  |    255 | lemma emeasure_not_MInf[simp]: "emeasure M A \<noteq> - \<infinity>"
 | 
|  |    256 |   using emeasure_nonneg[of M A] by auto
 | 
|  |    257 |   
 | 
|  |    258 | lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
 | 
|  |    259 |   by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
 | 
|  |    260 | 
 | 
|  |    261 | lemma suminf_emeasure:
 | 
|  |    262 |   "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
 | 
|  |    263 |   using countable_UN[of A UNIV M] emeasure_countably_additive[of M]
 | 
|  |    264 |   by (simp add: countably_additive_def)
 | 
|  |    265 | 
 | 
|  |    266 | lemma emeasure_additive: "additive (sets M) (emeasure M)"
 | 
|  |    267 |   by (metis countably_additive_additive emeasure_positive emeasure_countably_additive)
 | 
|  |    268 | 
 | 
|  |    269 | lemma plus_emeasure:
 | 
|  |    270 |   "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)"
 | 
|  |    271 |   using additiveD[OF emeasure_additive] ..
 | 
|  |    272 | 
 | 
|  |    273 | lemma setsum_emeasure:
 | 
|  |    274 |   "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
 | 
|  |    275 |     (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
 | 
|  |    276 |   by (metis additive_sum emeasure_positive emeasure_additive)
 | 
|  |    277 | 
 | 
|  |    278 | lemma emeasure_mono:
 | 
|  |    279 |   "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
 | 
|  |    280 |   by (metis additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets
 | 
|  |    281 |             emeasure_positive increasingD)
 | 
|  |    282 | 
 | 
|  |    283 | lemma emeasure_space:
 | 
|  |    284 |   "emeasure M A \<le> emeasure M (space M)"
 | 
|  |    285 |   by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets_into_space top)
 | 
|  |    286 | 
 | 
|  |    287 | lemma emeasure_compl:
 | 
|  |    288 |   assumes s: "s \<in> sets M" and fin: "emeasure M s \<noteq> \<infinity>"
 | 
|  |    289 |   shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
 | 
|  |    290 | proof -
 | 
|  |    291 |   from s have "0 \<le> emeasure M s" by auto
 | 
|  |    292 |   have "emeasure M (space M) = emeasure M (s \<union> (space M - s))" using s
 | 
|  |    293 |     by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
 | 
|  |    294 |   also have "... = emeasure M s + emeasure M (space M - s)"
 | 
|  |    295 |     by (rule plus_emeasure[symmetric]) (auto simp add: s)
 | 
|  |    296 |   finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" .
 | 
|  |    297 |   then show ?thesis
 | 
|  |    298 |     using fin `0 \<le> emeasure M s`
 | 
|  |    299 |     unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
 | 
|  |    300 | qed
 | 
|  |    301 | 
 | 
|  |    302 | lemma emeasure_Diff:
 | 
|  |    303 |   assumes finite: "emeasure M B \<noteq> \<infinity>"
 | 
|  |    304 |   and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
 | 
|  |    305 |   shows "emeasure M (A - B) = emeasure M A - emeasure M B"
 | 
|  |    306 | proof -
 | 
|  |    307 |   have "0 \<le> emeasure M B" using assms by auto
 | 
|  |    308 |   have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
 | 
|  |    309 |   then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
 | 
|  |    310 |   also have "\<dots> = emeasure M (A - B) + emeasure M B"
 | 
|  |    311 |     using measurable by (subst plus_emeasure[symmetric]) auto
 | 
|  |    312 |   finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
 | 
|  |    313 |     unfolding ereal_eq_minus_iff
 | 
|  |    314 |     using finite `0 \<le> emeasure M B` by auto
 | 
|  |    315 | qed
 | 
|  |    316 | 
 | 
|  |    317 | lemma emeasure_countable_increasing:
 | 
|  |    318 |   assumes A: "range A \<subseteq> sets M"
 | 
|  |    319 |       and A0: "A 0 = {}"
 | 
|  |    320 |       and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
 | 
|  |    321 |   shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
 | 
|  |    322 | proof -
 | 
|  |    323 |   { fix n
 | 
|  |    324 |     have "emeasure M (A n) = (\<Sum>i<n. emeasure M (A (Suc i) - A i))"
 | 
|  |    325 |       proof (induct n)
 | 
|  |    326 |         case 0 thus ?case by (auto simp add: A0)
 | 
|  |    327 |       next
 | 
|  |    328 |         case (Suc m)
 | 
|  |    329 |         have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
 | 
|  |    330 |           by (metis ASuc Un_Diff_cancel Un_absorb1)
 | 
|  |    331 |         hence "emeasure M (A (Suc m)) =
 | 
|  |    332 |                emeasure M (A m) + emeasure M (A (Suc m) - A m)"
 | 
|  |    333 |           by (subst plus_emeasure)
 | 
|  |    334 |              (auto simp add: emeasure_additive range_subsetD [OF A])
 | 
|  |    335 |         with Suc show ?case
 | 
|  |    336 |           by simp
 | 
|  |    337 |       qed }
 | 
|  |    338 |   note Meq = this
 | 
|  |    339 |   have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
 | 
|  |    340 |     proof (rule UN_finite2_eq [where k=1], simp)
 | 
|  |    341 |       fix i
 | 
|  |    342 |       show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
 | 
|  |    343 |         proof (induct i)
 | 
|  |    344 |           case 0 thus ?case by (simp add: A0)
 | 
|  |    345 |         next
 | 
|  |    346 |           case (Suc i)
 | 
|  |    347 |           thus ?case
 | 
|  |    348 |             by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
 | 
|  |    349 |         qed
 | 
|  |    350 |     qed
 | 
|  |    351 |   have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
 | 
|  |    352 |     by (metis A Diff range_subsetD)
 | 
|  |    353 |   have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
 | 
|  |    354 |     by (blast intro: range_subsetD [OF A])
 | 
|  |    355 |   have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = (\<Sum>i. emeasure M (A (Suc i) - A i))"
 | 
|  |    356 |     using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
 | 
|  |    357 |   also have "\<dots> = emeasure M (\<Union>i. A (Suc i) - A i)"
 | 
|  |    358 |     by (rule suminf_emeasure)
 | 
|  |    359 |        (auto simp add: disjoint_family_Suc ASuc A1 A2)
 | 
|  |    360 |   also have "... =  emeasure M (\<Union>i. A i)"
 | 
|  |    361 |     by (simp add: Aeq)
 | 
|  |    362 |   finally have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = emeasure M (\<Union>i. A i)" .
 | 
|  |    363 |   then show ?thesis by (auto simp add: Meq)
 | 
|  |    364 | qed
 | 
|  |    365 | 
 | 
|  |    366 | lemma SUP_emeasure_incseq:
 | 
|  |    367 |   assumes A: "range A \<subseteq> sets M" and "incseq A"
 | 
|  |    368 |   shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
 | 
|  |    369 | proof -
 | 
|  |    370 |   have *: "(SUP n. emeasure M (nat_case {} A (Suc n))) = (SUP n. emeasure M (nat_case {} A n))"
 | 
|  |    371 |     using A by (auto intro!: SUPR_eq exI split: nat.split)
 | 
|  |    372 |   have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
 | 
|  |    373 |     by (auto simp add: split: nat.splits)
 | 
|  |    374 |   have meq: "\<And>n. emeasure M (A n) = (emeasure M \<circ> nat_case {} A) (Suc n)"
 | 
|  |    375 |     by simp
 | 
|  |    376 |   have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. nat_case {} A i)"
 | 
|  |    377 |     using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
 | 
|  |    378 |     by (force split: nat.splits intro!: emeasure_countable_increasing)
 | 
|  |    379 |   also have "emeasure M (\<Union>i. nat_case {} A i) = emeasure M (\<Union>i. A i)"
 | 
|  |    380 |     by (simp add: ueq)
 | 
|  |    381 |   finally have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. A i)" .
 | 
|  |    382 |   thus ?thesis unfolding meq * comp_def .
 | 
|  |    383 | qed
 | 
|  |    384 | 
 | 
|  |    385 | lemma incseq_emeasure:
 | 
|  |    386 |   assumes "range B \<subseteq> sets M" "incseq B"
 | 
|  |    387 |   shows "incseq (\<lambda>i. emeasure M (B i))"
 | 
|  |    388 |   using assms by (auto simp: incseq_def intro!: emeasure_mono)
 | 
|  |    389 | 
 | 
|  |    390 | lemma Lim_emeasure_incseq:
 | 
|  |    391 |   assumes A: "range A \<subseteq> sets M" "incseq A"
 | 
|  |    392 |   shows "(\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
 | 
|  |    393 |   using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A]
 | 
|  |    394 |     SUP_emeasure_incseq[OF A] by simp
 | 
|  |    395 | 
 | 
|  |    396 | lemma decseq_emeasure:
 | 
|  |    397 |   assumes "range B \<subseteq> sets M" "decseq B"
 | 
|  |    398 |   shows "decseq (\<lambda>i. emeasure M (B i))"
 | 
|  |    399 |   using assms by (auto simp: decseq_def intro!: emeasure_mono)
 | 
|  |    400 | 
 | 
|  |    401 | lemma INF_emeasure_decseq:
 | 
|  |    402 |   assumes A: "range A \<subseteq> sets M" and "decseq A"
 | 
|  |    403 |   and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |    404 |   shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
 | 
|  |    405 | proof -
 | 
|  |    406 |   have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
 | 
|  |    407 |     using A by (auto intro!: emeasure_mono)
 | 
|  |    408 |   hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
 | 
|  |    409 | 
 | 
|  |    410 |   have A0: "0 \<le> emeasure M (A 0)" using A by auto
 | 
|  |    411 | 
 | 
|  |    412 |   have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))"
 | 
|  |    413 |     by (simp add: ereal_SUPR_uminus minus_ereal_def)
 | 
|  |    414 |   also have "\<dots> = (SUP n. emeasure M (A 0) - emeasure M (A n))"
 | 
|  |    415 |     unfolding minus_ereal_def using A0 assms
 | 
|  |    416 |     by (subst SUPR_ereal_add) (auto simp add: decseq_emeasure)
 | 
|  |    417 |   also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
 | 
|  |    418 |     using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto
 | 
|  |    419 |   also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
 | 
|  |    420 |   proof (rule SUP_emeasure_incseq)
 | 
|  |    421 |     show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
 | 
|  |    422 |       using A by auto
 | 
|  |    423 |     show "incseq (\<lambda>n. A 0 - A n)"
 | 
|  |    424 |       using `decseq A` by (auto simp add: incseq_def decseq_def)
 | 
|  |    425 |   qed
 | 
|  |    426 |   also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
 | 
|  |    427 |     using A finite * by (simp, subst emeasure_Diff) auto
 | 
|  |    428 |   finally show ?thesis
 | 
|  |    429 |     unfolding ereal_minus_eq_minus_iff using finite A0 by auto
 | 
|  |    430 | qed
 | 
|  |    431 | 
 | 
|  |    432 | lemma Lim_emeasure_decseq:
 | 
|  |    433 |   assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |    434 |   shows "(\<lambda>i. emeasure M (A i)) ----> emeasure M (\<Inter>i. A i)"
 | 
|  |    435 |   using LIMSEQ_ereal_INFI[OF decseq_emeasure, OF A]
 | 
|  |    436 |   using INF_emeasure_decseq[OF A fin] by simp
 | 
|  |    437 | 
 | 
|  |    438 | lemma emeasure_subadditive:
 | 
|  |    439 |   assumes measurable: "A \<in> sets M" "B \<in> sets M"
 | 
|  |    440 |   shows "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
 | 
|  |    441 | proof -
 | 
|  |    442 |   from plus_emeasure[of A M "B - A"]
 | 
|  |    443 |   have "emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)"
 | 
|  |    444 |     using assms by (simp add: Diff)
 | 
|  |    445 |   also have "\<dots> \<le> emeasure M A + emeasure M B"
 | 
|  |    446 |     using assms by (auto intro!: add_left_mono emeasure_mono)
 | 
|  |    447 |   finally show ?thesis .
 | 
|  |    448 | qed
 | 
|  |    449 | 
 | 
|  |    450 | lemma emeasure_subadditive_finite:
 | 
|  |    451 |   assumes "finite I" "A ` I \<subseteq> sets M"
 | 
|  |    452 |   shows "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
 | 
|  |    453 | using assms proof induct
 | 
|  |    454 |   case (insert i I)
 | 
|  |    455 |   then have "emeasure M (\<Union>i\<in>insert i I. A i) = emeasure M (A i \<union> (\<Union>i\<in>I. A i))"
 | 
|  |    456 |     by simp
 | 
|  |    457 |   also have "\<dots> \<le> emeasure M (A i) + emeasure M (\<Union>i\<in>I. A i)"
 | 
|  |    458 |     using insert by (intro emeasure_subadditive finite_UN) auto
 | 
|  |    459 |   also have "\<dots> \<le> emeasure M (A i) + (\<Sum>i\<in>I. emeasure M (A i))"
 | 
|  |    460 |     using insert by (intro add_mono) auto
 | 
|  |    461 |   also have "\<dots> = (\<Sum>i\<in>insert i I. emeasure M (A i))"
 | 
|  |    462 |     using insert by auto
 | 
|  |    463 |   finally show ?case .
 | 
|  |    464 | qed simp
 | 
|  |    465 | 
 | 
|  |    466 | lemma emeasure_subadditive_countably:
 | 
|  |    467 |   assumes "range f \<subseteq> sets M"
 | 
|  |    468 |   shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
 | 
|  |    469 | proof -
 | 
|  |    470 |   have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
 | 
|  |    471 |     unfolding UN_disjointed_eq ..
 | 
|  |    472 |   also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
 | 
|  |    473 |     using range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
 | 
|  |    474 |     by (simp add:  disjoint_family_disjointed comp_def)
 | 
|  |    475 |   also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
 | 
|  |    476 |     using range_disjointed_sets[OF assms] assms
 | 
|  |    477 |     by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset)
 | 
|  |    478 |   finally show ?thesis .
 | 
|  |    479 | qed
 | 
|  |    480 | 
 | 
|  |    481 | lemma emeasure_insert:
 | 
|  |    482 |   assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
 | 
|  |    483 |   shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
 | 
|  |    484 | proof -
 | 
|  |    485 |   have "{x} \<inter> A = {}" using `x \<notin> A` by auto
 | 
|  |    486 |   from plus_emeasure[OF sets this] show ?thesis by simp
 | 
|  |    487 | qed
 | 
|  |    488 | 
 | 
|  |    489 | lemma emeasure_eq_setsum_singleton:
 | 
|  |    490 |   assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
|  |    491 |   shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
 | 
|  |    492 |   using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
 | 
|  |    493 |   by (auto simp: disjoint_family_on_def subset_eq)
 | 
|  |    494 | 
 | 
|  |    495 | lemma setsum_emeasure_cover:
 | 
|  |    496 |   assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
 | 
|  |    497 |   assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
 | 
|  |    498 |   assumes disj: "disjoint_family_on B S"
 | 
|  |    499 |   shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
 | 
|  |    500 | proof -
 | 
|  |    501 |   have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
 | 
|  |    502 |   proof (rule setsum_emeasure)
 | 
|  |    503 |     show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
 | 
|  |    504 |       using `disjoint_family_on B S`
 | 
|  |    505 |       unfolding disjoint_family_on_def by auto
 | 
|  |    506 |   qed (insert assms, auto)
 | 
|  |    507 |   also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
 | 
|  |    508 |     using A by auto
 | 
|  |    509 |   finally show ?thesis by simp
 | 
|  |    510 | qed
 | 
|  |    511 | 
 | 
|  |    512 | lemma emeasure_eq_0:
 | 
|  |    513 |   "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
 | 
|  |    514 |   by (metis emeasure_mono emeasure_nonneg order_eq_iff)
 | 
|  |    515 | 
 | 
|  |    516 | lemma emeasure_UN_eq_0:
 | 
|  |    517 |   assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
 | 
|  |    518 |   shows "emeasure M (\<Union> i. N i) = 0"
 | 
|  |    519 | proof -
 | 
|  |    520 |   have "0 \<le> emeasure M (\<Union> i. N i)" using assms by auto
 | 
|  |    521 |   moreover have "emeasure M (\<Union> i. N i) \<le> 0"
 | 
|  |    522 |     using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
 | 
|  |    523 |   ultimately show ?thesis by simp
 | 
|  |    524 | qed
 | 
|  |    525 | 
 | 
|  |    526 | lemma measure_eqI_finite:
 | 
|  |    527 |   assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
 | 
|  |    528 |   assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
 | 
|  |    529 |   shows "M = N"
 | 
|  |    530 | proof (rule measure_eqI)
 | 
|  |    531 |   fix X assume "X \<in> sets M"
 | 
|  |    532 |   then have X: "X \<subseteq> A" by auto
 | 
|  |    533 |   then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
 | 
|  |    534 |     using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
 | 
|  |    535 |   also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
 | 
|  |    536 |     using X eq by (auto intro!: setsum_cong)
 | 
|  |    537 |   also have "\<dots> = emeasure N X"
 | 
|  |    538 |     using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
 | 
|  |    539 |   finally show "emeasure M X = emeasure N X" .
 | 
|  |    540 | qed simp
 | 
|  |    541 | 
 | 
|  |    542 | lemma measure_eqI_generator_eq:
 | 
|  |    543 |   fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
 | 
|  |    544 |   assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
 | 
|  |    545 |   and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
 | 
|  |    546 |   and M: "sets M = sigma_sets \<Omega> E"
 | 
|  |    547 |   and N: "sets N = sigma_sets \<Omega> E"
 | 
|  |    548 |   and A: "range A \<subseteq> E" "incseq A" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |    549 |   shows "M = N"
 | 
|  |    550 | proof -
 | 
|  |    551 |   let ?D = "\<lambda>F. {D. D \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)}"
 | 
|  |    552 |   interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
 | 
|  |    553 |   { fix F assume "F \<in> E" and "emeasure M F \<noteq> \<infinity>"
 | 
|  |    554 |     then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
 | 
|  |    555 |     have "emeasure N F \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` `F \<in> E` eq by simp
 | 
|  |    556 |     interpret D: dynkin_system \<Omega> "?D F"
 | 
|  |    557 |     proof (rule dynkin_systemI, simp_all)
 | 
|  |    558 |       fix A assume "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
 | 
|  |    559 |       then show "A \<subseteq> \<Omega>" using S.sets_into_space by auto
 | 
|  |    560 |     next
 | 
|  |    561 |       have "F \<inter> \<Omega> = F" using `F \<in> E` `E \<subseteq> Pow \<Omega>` by auto
 | 
|  |    562 |       then show "emeasure M (F \<inter> \<Omega>) = emeasure N (F \<inter> \<Omega>)"
 | 
|  |    563 |         using `F \<in> E` eq by (auto intro: sigma_sets_top)
 | 
|  |    564 |     next
 | 
|  |    565 |       fix A assume *: "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
 | 
|  |    566 |       then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
 | 
|  |    567 |         and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
 | 
|  |    568 |         using `F \<in> E` S.sets_into_space by auto
 | 
|  |    569 |       have "emeasure N (F \<inter> A) \<le> emeasure N F" by (auto intro!: emeasure_mono simp: M N)
 | 
|  |    570 |       then have "emeasure N (F \<inter> A) \<noteq> \<infinity>" using `emeasure N F \<noteq> \<infinity>` by auto
 | 
|  |    571 |       have "emeasure M (F \<inter> A) \<le> emeasure M F" by (auto intro!: emeasure_mono simp: M N)
 | 
|  |    572 |       then have "emeasure M (F \<inter> A) \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` by auto
 | 
|  |    573 |       then have "emeasure M (F \<inter> (\<Omega> - A)) = emeasure M F - emeasure M (F \<inter> A)" unfolding **
 | 
|  |    574 |         using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
 | 
|  |    575 |       also have "\<dots> = emeasure N F - emeasure N (F \<inter> A)" using eq `F \<in> E` * by simp
 | 
|  |    576 |       also have "\<dots> = emeasure N (F \<inter> (\<Omega> - A))" unfolding **
 | 
|  |    577 |         using `F \<inter> A \<in> sigma_sets \<Omega> E` `emeasure N (F \<inter> A) \<noteq> \<infinity>`
 | 
|  |    578 |         by (auto intro!: emeasure_Diff[symmetric] simp: M N)
 | 
|  |    579 |       finally show "\<Omega> - A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Omega> - A)) = emeasure N (F \<inter> (\<Omega> - A))"
 | 
|  |    580 |         using * by auto
 | 
|  |    581 |     next
 | 
|  |    582 |       fix A :: "nat \<Rightarrow> 'a set"
 | 
|  |    583 |       assume "disjoint_family A" "range A \<subseteq> {X \<in> sigma_sets \<Omega> E. emeasure M (F \<inter> X) = emeasure N (F \<inter> X)}"
 | 
|  |    584 |       then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sigma_sets \<Omega> E" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
 | 
|  |    585 |         "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. emeasure M (F \<inter> A i) = emeasure N (F \<inter> A i)" "range A \<subseteq> sigma_sets \<Omega> E"
 | 
|  |    586 |         by (auto simp: disjoint_family_on_def subset_eq)
 | 
|  |    587 |       then show "(\<Union>x. A x) \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Union>x. A x)) = emeasure N (F \<inter> (\<Union>x. A x))"
 | 
|  |    588 |         by (auto simp: M N suminf_emeasure[symmetric] simp del: UN_simps)
 | 
|  |    589 |     qed
 | 
|  |    590 |     have *: "sigma_sets \<Omega> E = ?D F"
 | 
|  |    591 |       using `F \<in> E` `Int_stable E`
 | 
|  |    592 |       by (intro D.dynkin_lemma) (auto simp add: Int_stable_def eq)
 | 
|  |    593 |     have "\<And>D. D \<in> sigma_sets \<Omega> E \<Longrightarrow> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
 | 
|  |    594 |       by (subst (asm) *) auto }
 | 
|  |    595 |   note * = this
 | 
|  |    596 |   show "M = N"
 | 
|  |    597 |   proof (rule measure_eqI)
 | 
|  |    598 |     show "sets M = sets N"
 | 
|  |    599 |       using M N by simp
 | 
|  |    600 |     fix X assume "X \<in> sets M"
 | 
|  |    601 |     then have "X \<in> sigma_sets \<Omega> E"
 | 
|  |    602 |       using M by simp
 | 
|  |    603 |     let ?A = "\<lambda>i. A i \<inter> X"
 | 
|  |    604 |     have "range ?A \<subseteq> sigma_sets \<Omega> E" "incseq ?A"
 | 
|  |    605 |       using A(1,2) `X \<in> sigma_sets \<Omega> E` by (auto simp: incseq_def)
 | 
|  |    606 |     moreover
 | 
|  |    607 |     { fix i have "emeasure M (?A i) = emeasure N (?A i)"
 | 
|  |    608 |         using *[of "A i" X] `X \<in> sigma_sets \<Omega> E` A finite by auto }
 | 
|  |    609 |     ultimately show "emeasure M X = emeasure N X"
 | 
|  |    610 |       using SUP_emeasure_incseq[of ?A M] SUP_emeasure_incseq[of ?A N] A(3) `X \<in> sigma_sets \<Omega> E`
 | 
|  |    611 |       by (auto simp: M N SUP_emeasure_incseq)
 | 
|  |    612 |   qed
 | 
|  |    613 | qed
 | 
|  |    614 | 
 | 
|  |    615 | lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
 | 
|  |    616 | proof (intro measure_eqI emeasure_measure_of_sigma)
 | 
|  |    617 |   show "sigma_algebra (space M) (sets M)" ..
 | 
|  |    618 |   show "positive (sets M) (emeasure M)"
 | 
|  |    619 |     by (simp add: positive_def emeasure_nonneg)
 | 
|  |    620 |   show "countably_additive (sets M) (emeasure M)"
 | 
|  |    621 |     by (simp add: emeasure_countably_additive)
 | 
|  |    622 | qed simp_all
 | 
|  |    623 | 
 | 
|  |    624 | section "@{text \<mu>}-null sets"
 | 
|  |    625 | 
 | 
|  |    626 | definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
 | 
|  |    627 |   "null_sets M = {N\<in>sets M. emeasure M N = 0}"
 | 
|  |    628 | 
 | 
|  |    629 | lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
 | 
|  |    630 |   by (simp add: null_sets_def)
 | 
|  |    631 | 
 | 
|  |    632 | lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
 | 
|  |    633 |   unfolding null_sets_def by simp
 | 
|  |    634 | 
 | 
|  |    635 | lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
 | 
|  |    636 |   unfolding null_sets_def by simp
 | 
|  |    637 | 
 | 
|  |    638 | interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
 | 
| 47762 |    639 | proof (rule ring_of_setsI)
 | 
| 47694 |    640 |   show "null_sets M \<subseteq> Pow (space M)"
 | 
|  |    641 |     using sets_into_space by auto
 | 
|  |    642 |   show "{} \<in> null_sets M"
 | 
|  |    643 |     by auto
 | 
|  |    644 |   fix A B assume sets: "A \<in> null_sets M" "B \<in> null_sets M"
 | 
|  |    645 |   then have "A \<in> sets M" "B \<in> sets M"
 | 
|  |    646 |     by auto
 | 
|  |    647 |   moreover then have "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
 | 
|  |    648 |     "emeasure M (A - B) \<le> emeasure M A"
 | 
|  |    649 |     by (auto intro!: emeasure_subadditive emeasure_mono)
 | 
|  |    650 |   moreover have "emeasure M B = 0" "emeasure M A = 0"
 | 
|  |    651 |     using sets by auto
 | 
|  |    652 |   ultimately show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
 | 
|  |    653 |     by (auto intro!: antisym)
 | 
|  |    654 | qed
 | 
|  |    655 | 
 | 
|  |    656 | lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))"
 | 
|  |    657 | proof -
 | 
|  |    658 |   have "(\<Union>i. N i) = (\<Union>i. (N \<circ> Countable.from_nat) i)"
 | 
|  |    659 |     unfolding SUP_def image_compose
 | 
|  |    660 |     unfolding surj_from_nat ..
 | 
|  |    661 |   then show ?thesis by simp
 | 
|  |    662 | qed
 | 
|  |    663 | 
 | 
|  |    664 | lemma null_sets_UN[intro]:
 | 
|  |    665 |   assumes "\<And>i::'i::countable. N i \<in> null_sets M"
 | 
|  |    666 |   shows "(\<Union>i. N i) \<in> null_sets M"
 | 
|  |    667 | proof (intro conjI CollectI null_setsI)
 | 
|  |    668 |   show "(\<Union>i. N i) \<in> sets M" using assms by auto
 | 
|  |    669 |   have "0 \<le> emeasure M (\<Union>i. N i)" by (rule emeasure_nonneg)
 | 
|  |    670 |   moreover have "emeasure M (\<Union>i. N i) \<le> (\<Sum>n. emeasure M (N (Countable.from_nat n)))"
 | 
|  |    671 |     unfolding UN_from_nat[of N]
 | 
|  |    672 |     using assms by (intro emeasure_subadditive_countably) auto
 | 
|  |    673 |   ultimately show "emeasure M (\<Union>i. N i) = 0"
 | 
|  |    674 |     using assms by (auto simp: null_setsD1)
 | 
|  |    675 | qed
 | 
|  |    676 | 
 | 
|  |    677 | lemma null_set_Int1:
 | 
|  |    678 |   assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
 | 
|  |    679 | proof (intro CollectI conjI null_setsI)
 | 
|  |    680 |   show "emeasure M (A \<inter> B) = 0" using assms
 | 
|  |    681 |     by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
 | 
|  |    682 | qed (insert assms, auto)
 | 
|  |    683 | 
 | 
|  |    684 | lemma null_set_Int2:
 | 
|  |    685 |   assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
 | 
|  |    686 |   using assms by (subst Int_commute) (rule null_set_Int1)
 | 
|  |    687 | 
 | 
|  |    688 | lemma emeasure_Diff_null_set:
 | 
|  |    689 |   assumes "B \<in> null_sets M" "A \<in> sets M"
 | 
|  |    690 |   shows "emeasure M (A - B) = emeasure M A"
 | 
|  |    691 | proof -
 | 
|  |    692 |   have *: "A - B = (A - (A \<inter> B))" by auto
 | 
|  |    693 |   have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
 | 
|  |    694 |   then show ?thesis
 | 
|  |    695 |     unfolding * using assms
 | 
|  |    696 |     by (subst emeasure_Diff) auto
 | 
|  |    697 | qed
 | 
|  |    698 | 
 | 
|  |    699 | lemma null_set_Diff:
 | 
|  |    700 |   assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
 | 
|  |    701 | proof (intro CollectI conjI null_setsI)
 | 
|  |    702 |   show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
 | 
|  |    703 | qed (insert assms, auto)
 | 
|  |    704 | 
 | 
|  |    705 | lemma emeasure_Un_null_set:
 | 
|  |    706 |   assumes "A \<in> sets M" "B \<in> null_sets M"
 | 
|  |    707 |   shows "emeasure M (A \<union> B) = emeasure M A"
 | 
|  |    708 | proof -
 | 
|  |    709 |   have *: "A \<union> B = A \<union> (B - A)" by auto
 | 
|  |    710 |   have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
 | 
|  |    711 |   then show ?thesis
 | 
|  |    712 |     unfolding * using assms
 | 
|  |    713 |     by (subst plus_emeasure[symmetric]) auto
 | 
|  |    714 | qed
 | 
|  |    715 | 
 | 
|  |    716 | section "Formalize almost everywhere"
 | 
|  |    717 | 
 | 
|  |    718 | definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
 | 
|  |    719 |   "ae_filter M = Abs_filter (\<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
 | 
|  |    720 | 
 | 
|  |    721 | abbreviation
 | 
|  |    722 |   almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
|  |    723 |   "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
 | 
|  |    724 | 
 | 
|  |    725 | syntax
 | 
|  |    726 |   "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
 | 
|  |    727 | 
 | 
|  |    728 | translations
 | 
|  |    729 |   "AE x in M. P" == "CONST almost_everywhere M (%x. P)"
 | 
|  |    730 | 
 | 
|  |    731 | lemma eventually_ae_filter:
 | 
|  |    732 |   fixes M P
 | 
|  |    733 |   defines [simp]: "F \<equiv> \<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N" 
 | 
|  |    734 |   shows "eventually P (ae_filter M) \<longleftrightarrow> F P"
 | 
|  |    735 |   unfolding ae_filter_def F_def[symmetric]
 | 
|  |    736 | proof (rule eventually_Abs_filter)
 | 
|  |    737 |   show "is_filter F"
 | 
|  |    738 |   proof
 | 
|  |    739 |     fix P Q assume "F P" "F Q"
 | 
|  |    740 |     then obtain N L where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N"
 | 
|  |    741 |       and L: "L \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> L"
 | 
|  |    742 |       by auto
 | 
|  |    743 |     then have "L \<union> N \<in> null_sets M" "{x \<in> space M. \<not> (P x \<and> Q x)} \<subseteq> L \<union> N" by auto
 | 
|  |    744 |     then show "F (\<lambda>x. P x \<and> Q x)" by auto
 | 
|  |    745 |   next
 | 
|  |    746 |     fix P Q assume "F P"
 | 
|  |    747 |     then obtain N where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N" by auto
 | 
|  |    748 |     moreover assume "\<forall>x. P x \<longrightarrow> Q x"
 | 
|  |    749 |     ultimately have "N \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> N" by auto
 | 
|  |    750 |     then show "F Q" by auto
 | 
|  |    751 |   qed auto
 | 
|  |    752 | qed
 | 
|  |    753 | 
 | 
|  |    754 | lemma AE_I':
 | 
|  |    755 |   "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
 | 
|  |    756 |   unfolding eventually_ae_filter by auto
 | 
|  |    757 | 
 | 
|  |    758 | lemma AE_iff_null:
 | 
|  |    759 |   assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
 | 
|  |    760 |   shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
 | 
|  |    761 | proof
 | 
|  |    762 |   assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
 | 
|  |    763 |     unfolding eventually_ae_filter by auto
 | 
|  |    764 |   have "0 \<le> emeasure M ?P" by auto
 | 
|  |    765 |   moreover have "emeasure M ?P \<le> emeasure M N"
 | 
|  |    766 |     using assms N(1,2) by (auto intro: emeasure_mono)
 | 
|  |    767 |   ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto
 | 
|  |    768 |   then show "?P \<in> null_sets M" using assms by auto
 | 
|  |    769 | next
 | 
|  |    770 |   assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
 | 
|  |    771 | qed
 | 
|  |    772 | 
 | 
|  |    773 | lemma AE_iff_null_sets:
 | 
|  |    774 |   "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
 | 
|  |    775 |   using Int_absorb1[OF sets_into_space, of N M]
 | 
|  |    776 |   by (subst AE_iff_null) (auto simp: Int_def[symmetric])
 | 
|  |    777 | 
 | 
| 47761 |    778 | lemma AE_not_in:
 | 
|  |    779 |   "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
 | 
|  |    780 |   by (metis AE_iff_null_sets null_setsD2)
 | 
|  |    781 | 
 | 
| 47694 |    782 | lemma AE_iff_measurable:
 | 
|  |    783 |   "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"
 | 
|  |    784 |   using AE_iff_null[of _ P] by auto
 | 
|  |    785 | 
 | 
|  |    786 | lemma AE_E[consumes 1]:
 | 
|  |    787 |   assumes "AE x in M. P x"
 | 
|  |    788 |   obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | 
|  |    789 |   using assms unfolding eventually_ae_filter by auto
 | 
|  |    790 | 
 | 
|  |    791 | lemma AE_E2:
 | 
|  |    792 |   assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
 | 
|  |    793 |   shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
 | 
|  |    794 | proof -
 | 
|  |    795 |   have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
 | 
|  |    796 |   with AE_iff_null[of M P] assms show ?thesis by auto
 | 
|  |    797 | qed
 | 
|  |    798 | 
 | 
|  |    799 | lemma AE_I:
 | 
|  |    800 |   assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
 | 
|  |    801 |   shows "AE x in M. P x"
 | 
|  |    802 |   using assms unfolding eventually_ae_filter by auto
 | 
|  |    803 | 
 | 
|  |    804 | lemma AE_mp[elim!]:
 | 
|  |    805 |   assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
 | 
|  |    806 |   shows "AE x in M. Q x"
 | 
|  |    807 | proof -
 | 
|  |    808 |   from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
 | 
|  |    809 |     and A: "A \<in> sets M" "emeasure M A = 0"
 | 
|  |    810 |     by (auto elim!: AE_E)
 | 
|  |    811 | 
 | 
|  |    812 |   from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
 | 
|  |    813 |     and B: "B \<in> sets M" "emeasure M B = 0"
 | 
|  |    814 |     by (auto elim!: AE_E)
 | 
|  |    815 | 
 | 
|  |    816 |   show ?thesis
 | 
|  |    817 |   proof (intro AE_I)
 | 
|  |    818 |     have "0 \<le> emeasure M (A \<union> B)" using A B by auto
 | 
|  |    819 |     moreover have "emeasure M (A \<union> B) \<le> 0"
 | 
|  |    820 |       using emeasure_subadditive[of A M B] A B by auto
 | 
|  |    821 |     ultimately show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" using A B by auto
 | 
|  |    822 |     show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
 | 
|  |    823 |       using P imp by auto
 | 
|  |    824 |   qed
 | 
|  |    825 | qed
 | 
|  |    826 | 
 | 
|  |    827 | (* depricated replace by laws about eventually *)
 | 
|  |    828 | lemma
 | 
|  |    829 |   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"
 | 
|  |    830 |     and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
 | 
|  |    831 |     and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
 | 
|  |    832 |     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"
 | 
|  |    833 |     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)"
 | 
|  |    834 |   by auto
 | 
|  |    835 | 
 | 
|  |    836 | lemma AE_impI:
 | 
|  |    837 |   "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
 | 
|  |    838 |   by (cases P) auto
 | 
|  |    839 | 
 | 
|  |    840 | lemma AE_measure:
 | 
|  |    841 |   assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
 | 
|  |    842 |   shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
 | 
|  |    843 | proof -
 | 
|  |    844 |   from AE_E[OF AE] guess N . note N = this
 | 
|  |    845 |   with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
 | 
|  |    846 |     by (intro emeasure_mono) auto
 | 
|  |    847 |   also have "\<dots> \<le> emeasure M ?P + emeasure M N"
 | 
|  |    848 |     using sets N by (intro emeasure_subadditive) auto
 | 
|  |    849 |   also have "\<dots> = emeasure M ?P" using N by simp
 | 
|  |    850 |   finally show "emeasure M ?P = emeasure M (space M)"
 | 
|  |    851 |     using emeasure_space[of M "?P"] by auto
 | 
|  |    852 | qed
 | 
|  |    853 | 
 | 
|  |    854 | lemma AE_space: "AE x in M. x \<in> space M"
 | 
|  |    855 |   by (rule AE_I[where N="{}"]) auto
 | 
|  |    856 | 
 | 
|  |    857 | lemma AE_I2[simp, intro]:
 | 
|  |    858 |   "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
 | 
|  |    859 |   using AE_space by force
 | 
|  |    860 | 
 | 
|  |    861 | lemma AE_Ball_mp:
 | 
|  |    862 |   "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
 | 
|  |    863 |   by auto
 | 
|  |    864 | 
 | 
|  |    865 | lemma AE_cong[cong]:
 | 
|  |    866 |   "(\<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)"
 | 
|  |    867 |   by auto
 | 
|  |    868 | 
 | 
|  |    869 | lemma AE_all_countable:
 | 
|  |    870 |   "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
 | 
|  |    871 | proof
 | 
|  |    872 |   assume "\<forall>i. AE x in M. P i x"
 | 
|  |    873 |   from this[unfolded eventually_ae_filter Bex_def, THEN choice]
 | 
|  |    874 |   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
 | 
|  |    875 |   have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
 | 
|  |    876 |   also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
 | 
|  |    877 |   finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
 | 
|  |    878 |   moreover from N have "(\<Union>i. N i) \<in> null_sets M"
 | 
|  |    879 |     by (intro null_sets_UN) auto
 | 
|  |    880 |   ultimately show "AE x in M. \<forall>i. P i x"
 | 
|  |    881 |     unfolding eventually_ae_filter by auto
 | 
|  |    882 | qed auto
 | 
|  |    883 | 
 | 
|  |    884 | lemma AE_finite_all:
 | 
|  |    885 |   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)"
 | 
|  |    886 |   using f by induct auto
 | 
|  |    887 | 
 | 
|  |    888 | lemma AE_finite_allI:
 | 
|  |    889 |   assumes "finite S"
 | 
|  |    890 |   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"
 | 
|  |    891 |   using AE_finite_all[OF `finite S`] by auto
 | 
|  |    892 | 
 | 
|  |    893 | lemma emeasure_mono_AE:
 | 
|  |    894 |   assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
 | 
|  |    895 |     and B: "B \<in> sets M"
 | 
|  |    896 |   shows "emeasure M A \<le> emeasure M B"
 | 
|  |    897 | proof cases
 | 
|  |    898 |   assume A: "A \<in> sets M"
 | 
|  |    899 |   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"
 | 
|  |    900 |     by (auto simp: eventually_ae_filter)
 | 
|  |    901 |   have "emeasure M A = emeasure M (A - N)"
 | 
|  |    902 |     using N A by (subst emeasure_Diff_null_set) auto
 | 
|  |    903 |   also have "emeasure M (A - N) \<le> emeasure M (B - N)"
 | 
|  |    904 |     using N A B sets_into_space by (auto intro!: emeasure_mono)
 | 
|  |    905 |   also have "emeasure M (B - N) = emeasure M B"
 | 
|  |    906 |     using N B by (subst emeasure_Diff_null_set) auto
 | 
|  |    907 |   finally show ?thesis .
 | 
|  |    908 | qed (simp add: emeasure_nonneg emeasure_notin_sets)
 | 
|  |    909 | 
 | 
|  |    910 | lemma emeasure_eq_AE:
 | 
|  |    911 |   assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
 | 
|  |    912 |   assumes A: "A \<in> sets M" and B: "B \<in> sets M"
 | 
|  |    913 |   shows "emeasure M A = emeasure M B"
 | 
|  |    914 |   using assms by (safe intro!: antisym emeasure_mono_AE) auto
 | 
|  |    915 | 
 | 
|  |    916 | section {* @{text \<sigma>}-finite Measures *}
 | 
|  |    917 | 
 | 
|  |    918 | locale sigma_finite_measure =
 | 
|  |    919 |   fixes M :: "'a measure"
 | 
|  |    920 |   assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set.
 | 
|  |    921 |     range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
 | 
|  |    922 | 
 | 
|  |    923 | lemma (in sigma_finite_measure) sigma_finite_disjoint:
 | 
|  |    924 |   obtains A :: "nat \<Rightarrow> 'a set"
 | 
|  |    925 |   where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
 | 
|  |    926 | proof atomize_elim
 | 
|  |    927 |   case goal1
 | 
|  |    928 |   obtain A :: "nat \<Rightarrow> 'a set" where
 | 
|  |    929 |     range: "range A \<subseteq> sets M" and
 | 
|  |    930 |     space: "(\<Union>i. A i) = space M" and
 | 
|  |    931 |     measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |    932 |     using sigma_finite by auto
 | 
|  |    933 |   note range' = range_disjointed_sets[OF range] range
 | 
|  |    934 |   { fix i
 | 
|  |    935 |     have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
 | 
|  |    936 |       using range' disjointed_subset[of A i] by (auto intro!: emeasure_mono)
 | 
|  |    937 |     then have "emeasure M (disjointed A i) \<noteq> \<infinity>"
 | 
|  |    938 |       using measure[of i] by auto }
 | 
|  |    939 |   with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
 | 
|  |    940 |   show ?case by (auto intro!: exI[of _ "disjointed A"])
 | 
|  |    941 | qed
 | 
|  |    942 | 
 | 
|  |    943 | lemma (in sigma_finite_measure) sigma_finite_incseq:
 | 
|  |    944 |   obtains A :: "nat \<Rightarrow> 'a set"
 | 
|  |    945 |   where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
 | 
|  |    946 | proof atomize_elim
 | 
|  |    947 |   case goal1
 | 
|  |    948 |   obtain F :: "nat \<Rightarrow> 'a set" where
 | 
|  |    949 |     F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
 | 
|  |    950 |     using sigma_finite by auto
 | 
|  |    951 |   then show ?case
 | 
|  |    952 |   proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI)
 | 
|  |    953 |     from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
 | 
|  |    954 |     then show "(\<Union>n. \<Union> i\<le>n. F i) = space M"
 | 
|  |    955 |       using F by fastforce
 | 
|  |    956 |   next
 | 
|  |    957 |     fix n
 | 
|  |    958 |     have "emeasure M (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))" using F
 | 
|  |    959 |       by (auto intro!: emeasure_subadditive_finite)
 | 
|  |    960 |     also have "\<dots> < \<infinity>"
 | 
|  |    961 |       using F by (auto simp: setsum_Pinfty)
 | 
|  |    962 |     finally show "emeasure M (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
 | 
|  |    963 |   qed (force simp: incseq_def)+
 | 
|  |    964 | qed
 | 
|  |    965 | 
 | 
|  |    966 | section {* Measure space induced by distribution of @{const measurable}-functions *}
 | 
|  |    967 | 
 | 
|  |    968 | definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
 | 
|  |    969 |   "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
 | 
|  |    970 | 
 | 
|  |    971 | lemma
 | 
|  |    972 |   shows sets_distr[simp]: "sets (distr M N f) = sets N"
 | 
|  |    973 |     and space_distr[simp]: "space (distr M N f) = space N"
 | 
|  |    974 |   by (auto simp: distr_def)
 | 
|  |    975 | 
 | 
|  |    976 | lemma
 | 
|  |    977 |   shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
 | 
|  |    978 |     and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
 | 
|  |    979 |   by (auto simp: measurable_def)
 | 
|  |    980 | 
 | 
|  |    981 | lemma emeasure_distr:
 | 
|  |    982 |   fixes f :: "'a \<Rightarrow> 'b"
 | 
|  |    983 |   assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
 | 
|  |    984 |   shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
 | 
|  |    985 |   unfolding distr_def
 | 
|  |    986 | proof (rule emeasure_measure_of_sigma)
 | 
|  |    987 |   show "positive (sets N) ?\<mu>"
 | 
|  |    988 |     by (auto simp: positive_def)
 | 
|  |    989 | 
 | 
|  |    990 |   show "countably_additive (sets N) ?\<mu>"
 | 
|  |    991 |   proof (intro countably_additiveI)
 | 
|  |    992 |     fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
 | 
|  |    993 |     then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
 | 
|  |    994 |     then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
 | 
|  |    995 |       using f by (auto simp: measurable_def)
 | 
|  |    996 |     moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
 | 
|  |    997 |       using * by blast
 | 
|  |    998 |     moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
 | 
|  |    999 |       using `disjoint_family A` by (auto simp: disjoint_family_on_def)
 | 
|  |   1000 |     ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
 | 
|  |   1001 |       using suminf_emeasure[OF _ **] A f
 | 
|  |   1002 |       by (auto simp: comp_def vimage_UN)
 | 
|  |   1003 |   qed
 | 
|  |   1004 |   show "sigma_algebra (space N) (sets N)" ..
 | 
|  |   1005 | qed fact
 | 
|  |   1006 | 
 | 
|  |   1007 | lemma AE_distrD:
 | 
|  |   1008 |   assumes f: "f \<in> measurable M M'"
 | 
|  |   1009 |     and AE: "AE x in distr M M' f. P x"
 | 
|  |   1010 |   shows "AE x in M. P (f x)"
 | 
|  |   1011 | proof -
 | 
|  |   1012 |   from AE[THEN AE_E] guess N .
 | 
|  |   1013 |   with f show ?thesis
 | 
|  |   1014 |     unfolding eventually_ae_filter
 | 
|  |   1015 |     by (intro bexI[of _ "f -` N \<inter> space M"])
 | 
|  |   1016 |        (auto simp: emeasure_distr measurable_def)
 | 
|  |   1017 | qed
 | 
|  |   1018 | 
 | 
|  |   1019 | lemma null_sets_distr_iff:
 | 
|  |   1020 |   "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"
 | 
|  |   1021 |   by (auto simp add: null_sets_def emeasure_distr measurable_sets)
 | 
|  |   1022 | 
 | 
|  |   1023 | lemma distr_distr:
 | 
|  |   1024 |   assumes f: "g \<in> measurable N L" and g: "f \<in> measurable M N"
 | 
|  |   1025 |   shows "distr (distr M N f) L g = distr M L (g \<circ> f)" (is "?L = ?R")
 | 
|  |   1026 |   using measurable_comp[OF g f] f g
 | 
|  |   1027 |   by (auto simp add: emeasure_distr measurable_sets measurable_space
 | 
|  |   1028 |            intro!: arg_cong[where f="emeasure M"] measure_eqI)
 | 
|  |   1029 | 
 | 
|  |   1030 | section {* Real measure values *}
 | 
|  |   1031 | 
 | 
|  |   1032 | lemma measure_nonneg: "0 \<le> measure M A"
 | 
|  |   1033 |   using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos)
 | 
|  |   1034 | 
 | 
|  |   1035 | lemma measure_empty[simp]: "measure M {} = 0"
 | 
|  |   1036 |   unfolding measure_def by simp
 | 
|  |   1037 | 
 | 
|  |   1038 | lemma emeasure_eq_ereal_measure:
 | 
|  |   1039 |   "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M A = ereal (measure M A)"
 | 
|  |   1040 |   using emeasure_nonneg[of M A]
 | 
|  |   1041 |   by (cases "emeasure M A") (auto simp: measure_def)
 | 
|  |   1042 | 
 | 
|  |   1043 | lemma measure_Union:
 | 
|  |   1044 |   assumes finite: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
 | 
|  |   1045 |   and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
 | 
|  |   1046 |   shows "measure M (A \<union> B) = measure M A + measure M B"
 | 
|  |   1047 |   unfolding measure_def
 | 
|  |   1048 |   using plus_emeasure[OF measurable, symmetric] finite
 | 
|  |   1049 |   by (simp add: emeasure_eq_ereal_measure)
 | 
|  |   1050 | 
 | 
|  |   1051 | lemma measure_finite_Union:
 | 
|  |   1052 |   assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
 | 
|  |   1053 |   assumes finite: "\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |   1054 |   shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
 | 
|  |   1055 |   unfolding measure_def
 | 
|  |   1056 |   using setsum_emeasure[OF measurable, symmetric] finite
 | 
|  |   1057 |   by (simp add: emeasure_eq_ereal_measure)
 | 
|  |   1058 | 
 | 
|  |   1059 | lemma measure_Diff:
 | 
|  |   1060 |   assumes finite: "emeasure M A \<noteq> \<infinity>"
 | 
|  |   1061 |   and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
 | 
|  |   1062 |   shows "measure M (A - B) = measure M A - measure M B"
 | 
|  |   1063 | proof -
 | 
|  |   1064 |   have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
 | 
|  |   1065 |     using measurable by (auto intro!: emeasure_mono)
 | 
|  |   1066 |   hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
 | 
|  |   1067 |     using measurable finite by (rule_tac measure_Union) auto
 | 
|  |   1068 |   thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
 | 
|  |   1069 | qed
 | 
|  |   1070 | 
 | 
|  |   1071 | lemma measure_UNION:
 | 
|  |   1072 |   assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
 | 
|  |   1073 |   assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
 | 
|  |   1074 |   shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
 | 
|  |   1075 | proof -
 | 
|  |   1076 |   from summable_sums[OF summable_ereal_pos, of "\<lambda>i. emeasure M (A i)"]
 | 
|  |   1077 |        suminf_emeasure[OF measurable] emeasure_nonneg[of M]
 | 
|  |   1078 |   have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" by simp
 | 
|  |   1079 |   moreover
 | 
|  |   1080 |   { fix i
 | 
|  |   1081 |     have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
 | 
|  |   1082 |       using measurable by (auto intro!: emeasure_mono)
 | 
|  |   1083 |     then have "emeasure M (A i) = ereal ((measure M (A i)))"
 | 
|  |   1084 |       using finite by (intro emeasure_eq_ereal_measure) auto }
 | 
|  |   1085 |   ultimately show ?thesis using finite
 | 
|  |   1086 |     unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure)
 | 
|  |   1087 | qed
 | 
|  |   1088 | 
 | 
|  |   1089 | lemma measure_subadditive:
 | 
|  |   1090 |   assumes measurable: "A \<in> sets M" "B \<in> sets M"
 | 
|  |   1091 |   and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
 | 
|  |   1092 |   shows "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
 | 
|  |   1093 | proof -
 | 
|  |   1094 |   have "emeasure M (A \<union> B) \<noteq> \<infinity>"
 | 
|  |   1095 |     using emeasure_subadditive[OF measurable] fin by auto
 | 
|  |   1096 |   then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
 | 
|  |   1097 |     using emeasure_subadditive[OF measurable] fin
 | 
|  |   1098 |     by (auto simp: emeasure_eq_ereal_measure)
 | 
|  |   1099 | qed
 | 
|  |   1100 | 
 | 
|  |   1101 | lemma measure_subadditive_finite:
 | 
|  |   1102 |   assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |   1103 |   shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
 | 
|  |   1104 | proof -
 | 
|  |   1105 |   { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
 | 
|  |   1106 |       using emeasure_subadditive_finite[OF A] .
 | 
|  |   1107 |     also have "\<dots> < \<infinity>"
 | 
|  |   1108 |       using fin by (simp add: setsum_Pinfty)
 | 
|  |   1109 |     finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> \<infinity>" by simp }
 | 
|  |   1110 |   then show ?thesis
 | 
|  |   1111 |     using emeasure_subadditive_finite[OF A] fin
 | 
|  |   1112 |     unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
 | 
|  |   1113 | qed
 | 
|  |   1114 | 
 | 
|  |   1115 | lemma measure_subadditive_countably:
 | 
|  |   1116 |   assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
 | 
|  |   1117 |   shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
 | 
|  |   1118 | proof -
 | 
|  |   1119 |   from emeasure_nonneg fin have "\<And>i. emeasure M (A i) \<noteq> \<infinity>" by (rule suminf_PInfty)
 | 
|  |   1120 |   moreover
 | 
|  |   1121 |   { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
 | 
|  |   1122 |       using emeasure_subadditive_countably[OF A] .
 | 
|  |   1123 |     also have "\<dots> < \<infinity>"
 | 
|  |   1124 |       using fin by simp
 | 
|  |   1125 |     finally have "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" by simp }
 | 
|  |   1126 |   ultimately  show ?thesis
 | 
|  |   1127 |     using emeasure_subadditive_countably[OF A] fin
 | 
|  |   1128 |     unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
 | 
|  |   1129 | qed
 | 
|  |   1130 | 
 | 
|  |   1131 | lemma measure_eq_setsum_singleton:
 | 
|  |   1132 |   assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
|  |   1133 |   and fin: "\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>"
 | 
|  |   1134 |   shows "(measure M S) = (\<Sum>x\<in>S. (measure M {x}))"
 | 
|  |   1135 |   unfolding measure_def
 | 
|  |   1136 |   using emeasure_eq_setsum_singleton[OF S] fin
 | 
|  |   1137 |   by simp (simp add: emeasure_eq_ereal_measure)
 | 
|  |   1138 | 
 | 
|  |   1139 | lemma Lim_measure_incseq:
 | 
|  |   1140 |   assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
 | 
|  |   1141 |   shows "(\<lambda>i. (measure M (A i))) ----> (measure M (\<Union>i. A i))"
 | 
|  |   1142 | proof -
 | 
|  |   1143 |   have "ereal ((measure M (\<Union>i. A i))) = emeasure M (\<Union>i. A i)"
 | 
|  |   1144 |     using fin by (auto simp: emeasure_eq_ereal_measure)
 | 
|  |   1145 |   then show ?thesis
 | 
|  |   1146 |     using Lim_emeasure_incseq[OF A]
 | 
|  |   1147 |     unfolding measure_def
 | 
|  |   1148 |     by (intro lim_real_of_ereal) simp
 | 
|  |   1149 | qed
 | 
|  |   1150 | 
 | 
|  |   1151 | lemma Lim_measure_decseq:
 | 
|  |   1152 |   assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
 | 
|  |   1153 |   shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
 | 
|  |   1154 | proof -
 | 
|  |   1155 |   have "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
 | 
|  |   1156 |     using A by (auto intro!: emeasure_mono)
 | 
|  |   1157 |   also have "\<dots> < \<infinity>"
 | 
|  |   1158 |     using fin[of 0] by auto
 | 
|  |   1159 |   finally have "ereal ((measure M (\<Inter>i. A i))) = emeasure M (\<Inter>i. A i)"
 | 
|  |   1160 |     by (auto simp: emeasure_eq_ereal_measure)
 | 
|  |   1161 |   then show ?thesis
 | 
|  |   1162 |     unfolding measure_def
 | 
|  |   1163 |     using Lim_emeasure_decseq[OF A fin]
 | 
|  |   1164 |     by (intro lim_real_of_ereal) simp
 | 
|  |   1165 | qed
 | 
|  |   1166 | 
 | 
|  |   1167 | section {* Measure spaces with @{term "emeasure M (space M) < \<infinity>"} *}
 | 
|  |   1168 | 
 | 
|  |   1169 | locale finite_measure = sigma_finite_measure M for M +
 | 
|  |   1170 |   assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
 | 
|  |   1171 | 
 | 
|  |   1172 | lemma finite_measureI[Pure.intro!]:
 | 
|  |   1173 |   assumes *: "emeasure M (space M) \<noteq> \<infinity>"
 | 
|  |   1174 |   shows "finite_measure M"
 | 
|  |   1175 | proof
 | 
|  |   1176 |   show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
 | 
|  |   1177 |     using * by (auto intro!: exI[of _ "\<lambda>_. space M"])
 | 
|  |   1178 | qed fact
 | 
|  |   1179 | 
 | 
|  |   1180 | lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> \<infinity>"
 | 
|  |   1181 |   using finite_emeasure_space emeasure_space[of M A] by auto
 | 
|  |   1182 | 
 | 
|  |   1183 | lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)"
 | 
|  |   1184 |   unfolding measure_def by (simp add: emeasure_eq_ereal_measure)
 | 
|  |   1185 | 
 | 
|  |   1186 | lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ereal r"
 | 
|  |   1187 |   using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto
 | 
|  |   1188 | 
 | 
|  |   1189 | lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
 | 
|  |   1190 |   using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
 | 
|  |   1191 | 
 | 
|  |   1192 | lemma (in finite_measure) finite_measure_Diff:
 | 
|  |   1193 |   assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
 | 
|  |   1194 |   shows "measure M (A - B) = measure M A - measure M B"
 | 
|  |   1195 |   using measure_Diff[OF _ assms] by simp
 | 
|  |   1196 | 
 | 
|  |   1197 | lemma (in finite_measure) finite_measure_Union:
 | 
|  |   1198 |   assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
 | 
|  |   1199 |   shows "measure M (A \<union> B) = measure M A + measure M B"
 | 
|  |   1200 |   using measure_Union[OF _ _ assms] by simp
 | 
|  |   1201 | 
 | 
|  |   1202 | lemma (in finite_measure) finite_measure_finite_Union:
 | 
|  |   1203 |   assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
 | 
|  |   1204 |   shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
 | 
|  |   1205 |   using measure_finite_Union[OF assms] by simp
 | 
|  |   1206 | 
 | 
|  |   1207 | lemma (in finite_measure) finite_measure_UNION:
 | 
|  |   1208 |   assumes A: "range A \<subseteq> sets M" "disjoint_family A"
 | 
|  |   1209 |   shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
 | 
|  |   1210 |   using measure_UNION[OF A] by simp
 | 
|  |   1211 | 
 | 
|  |   1212 | lemma (in finite_measure) finite_measure_mono:
 | 
|  |   1213 |   assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
 | 
|  |   1214 |   using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
 | 
|  |   1215 | 
 | 
|  |   1216 | lemma (in finite_measure) finite_measure_subadditive:
 | 
|  |   1217 |   assumes m: "A \<in> sets M" "B \<in> sets M"
 | 
|  |   1218 |   shows "measure M (A \<union> B) \<le> measure M A + measure M B"
 | 
|  |   1219 |   using measure_subadditive[OF m] by simp
 | 
|  |   1220 | 
 | 
|  |   1221 | lemma (in finite_measure) finite_measure_subadditive_finite:
 | 
|  |   1222 |   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))"
 | 
|  |   1223 |   using measure_subadditive_finite[OF assms] by simp
 | 
|  |   1224 | 
 | 
|  |   1225 | lemma (in finite_measure) finite_measure_subadditive_countably:
 | 
|  |   1226 |   assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
 | 
|  |   1227 |   shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
 | 
|  |   1228 | proof -
 | 
|  |   1229 |   from `summable (\<lambda>i. measure M (A i))`
 | 
|  |   1230 |   have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>i. measure M (A i))"
 | 
|  |   1231 |     by (simp add: sums_ereal) (rule summable_sums)
 | 
|  |   1232 |   from sums_unique[OF this, symmetric]
 | 
|  |   1233 |        measure_subadditive_countably[OF A]
 | 
|  |   1234 |   show ?thesis by (simp add: emeasure_eq_measure)
 | 
|  |   1235 | qed
 | 
|  |   1236 | 
 | 
|  |   1237 | lemma (in finite_measure) finite_measure_eq_setsum_singleton:
 | 
|  |   1238 |   assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
|  |   1239 |   shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
 | 
|  |   1240 |   using measure_eq_setsum_singleton[OF assms] by simp
 | 
|  |   1241 | 
 | 
|  |   1242 | lemma (in finite_measure) finite_Lim_measure_incseq:
 | 
|  |   1243 |   assumes A: "range A \<subseteq> sets M" "incseq A"
 | 
|  |   1244 |   shows "(\<lambda>i. measure M (A i)) ----> measure M (\<Union>i. A i)"
 | 
|  |   1245 |   using Lim_measure_incseq[OF A] by simp
 | 
|  |   1246 | 
 | 
|  |   1247 | lemma (in finite_measure) finite_Lim_measure_decseq:
 | 
|  |   1248 |   assumes A: "range A \<subseteq> sets M" "decseq A"
 | 
|  |   1249 |   shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
 | 
|  |   1250 |   using Lim_measure_decseq[OF A] by simp
 | 
|  |   1251 | 
 | 
|  |   1252 | lemma (in finite_measure) finite_measure_compl:
 | 
|  |   1253 |   assumes S: "S \<in> sets M"
 | 
|  |   1254 |   shows "measure M (space M - S) = measure M (space M) - measure M S"
 | 
|  |   1255 |   using measure_Diff[OF _ top S sets_into_space] S by simp
 | 
|  |   1256 | 
 | 
|  |   1257 | lemma (in finite_measure) finite_measure_mono_AE:
 | 
|  |   1258 |   assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
 | 
|  |   1259 |   shows "measure M A \<le> measure M B"
 | 
|  |   1260 |   using assms emeasure_mono_AE[OF imp B]
 | 
|  |   1261 |   by (simp add: emeasure_eq_measure)
 | 
|  |   1262 | 
 | 
|  |   1263 | lemma (in finite_measure) finite_measure_eq_AE:
 | 
|  |   1264 |   assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
 | 
|  |   1265 |   assumes A: "A \<in> sets M" and B: "B \<in> sets M"
 | 
|  |   1266 |   shows "measure M A = measure M B"
 | 
|  |   1267 |   using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
 | 
|  |   1268 | 
 | 
|  |   1269 | section {* Counting space *}
 | 
|  |   1270 | 
 | 
|  |   1271 | definition count_space :: "'a set \<Rightarrow> 'a measure" where
 | 
|  |   1272 |   "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"
 | 
|  |   1273 | 
 | 
|  |   1274 | lemma 
 | 
|  |   1275 |   shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
 | 
|  |   1276 |     and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
 | 
|  |   1277 |   using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
 | 
|  |   1278 |   by (auto simp: count_space_def)
 | 
|  |   1279 | 
 | 
|  |   1280 | lemma measurable_count_space_eq1[simp]:
 | 
|  |   1281 |   "f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"
 | 
|  |   1282 |  unfolding measurable_def by simp
 | 
|  |   1283 | 
 | 
|  |   1284 | lemma measurable_count_space_eq2[simp]:
 | 
|  |   1285 |   assumes "finite A"
 | 
|  |   1286 |   shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
 | 
|  |   1287 | proof -
 | 
|  |   1288 |   { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
 | 
|  |   1289 |     with `finite A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "finite X"
 | 
|  |   1290 |       by (auto dest: finite_subset)
 | 
|  |   1291 |     moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
 | 
|  |   1292 |     ultimately have "f -` X \<inter> space M \<in> sets M"
 | 
|  |   1293 |       using `X \<subseteq> A` by (auto intro!: finite_UN simp del: UN_simps) }
 | 
|  |   1294 |   then show ?thesis
 | 
|  |   1295 |     unfolding measurable_def by auto
 | 
|  |   1296 | qed
 | 
|  |   1297 | 
 | 
|  |   1298 | lemma emeasure_count_space:
 | 
|  |   1299 |   assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \<infinity>)"
 | 
|  |   1300 |     (is "_ = ?M X")
 | 
|  |   1301 |   unfolding count_space_def
 | 
|  |   1302 | proof (rule emeasure_measure_of_sigma)
 | 
|  |   1303 |   show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
 | 
|  |   1304 | 
 | 
|  |   1305 |   show "positive (Pow A) ?M"
 | 
|  |   1306 |     by (auto simp: positive_def)
 | 
|  |   1307 | 
 | 
|  |   1308 |   show "countably_additive (Pow A) ?M"
 | 
|  |   1309 |   proof (unfold countably_additive_def, safe)
 | 
|  |   1310 |       fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F"
 | 
|  |   1311 |       show "(\<Sum>i. ?M (F i)) = ?M (\<Union>i. F i)"
 | 
|  |   1312 |       proof cases
 | 
|  |   1313 |         assume "\<forall>i. finite (F i)"
 | 
|  |   1314 |         then have finite_F: "\<And>i. finite (F i)" by auto
 | 
|  |   1315 |         have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
 | 
|  |   1316 |         from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
 | 
|  |   1317 | 
 | 
|  |   1318 |         have inj_f: "inj_on f {i. F i \<noteq> {}}"
 | 
|  |   1319 |         proof (rule inj_onI, simp)
 | 
|  |   1320 |           fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
 | 
|  |   1321 |           then have "f i \<in> F i" "f j \<in> F j" using f by force+
 | 
|  |   1322 |           with disj * show "i = j" by (auto simp: disjoint_family_on_def)
 | 
|  |   1323 |         qed
 | 
|  |   1324 |         have fin_eq: "finite (\<Union>i. F i) \<longleftrightarrow> finite {i. F i \<noteq> {}}"
 | 
|  |   1325 |         proof
 | 
|  |   1326 |           assume "finite (\<Union>i. F i)"
 | 
|  |   1327 |           show "finite {i. F i \<noteq> {}}"
 | 
|  |   1328 |           proof (rule finite_imageD)
 | 
|  |   1329 |             from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
 | 
|  |   1330 |             then show "finite (f`{i. F i \<noteq> {}})"
 | 
|  |   1331 |               by (rule finite_subset) fact
 | 
|  |   1332 |           qed fact
 | 
|  |   1333 |         next
 | 
|  |   1334 |           assume "finite {i. F i \<noteq> {}}"
 | 
|  |   1335 |           with finite_F have "finite (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
 | 
|  |   1336 |             by auto
 | 
|  |   1337 |           also have "(\<Union>i\<in>{i. F i \<noteq> {}}. F i) = (\<Union>i. F i)"
 | 
|  |   1338 |             by auto
 | 
|  |   1339 |           finally show "finite (\<Union>i. F i)" .
 | 
|  |   1340 |         qed
 | 
|  |   1341 |         
 | 
|  |   1342 |         show ?thesis
 | 
|  |   1343 |         proof cases
 | 
|  |   1344 |           assume *: "finite (\<Union>i. F i)"
 | 
|  |   1345 |           with finite_F have "finite {i. ?M (F i) \<noteq> 0} "
 | 
|  |   1346 |             by (simp add: fin_eq)
 | 
|  |   1347 |           then have "(\<Sum>i. ?M (F i)) = (\<Sum>i | ?M (F i) \<noteq> 0. ?M (F i))"
 | 
| 47761 |   1348 |             by (rule suminf_finite) auto
 | 
| 47694 |   1349 |           also have "\<dots> = ereal (\<Sum>i | F i \<noteq> {}. card (F i))"
 | 
|  |   1350 |             using finite_F by simp
 | 
|  |   1351 |           also have "\<dots> = ereal (card (\<Union>i \<in> {i. F i \<noteq> {}}. F i))"
 | 
|  |   1352 |             using * finite_F disj
 | 
|  |   1353 |             by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def fin_eq)
 | 
|  |   1354 |           also have "\<dots> = ?M (\<Union>i. F i)"
 | 
|  |   1355 |             using * by (auto intro!: arg_cong[where f=card])
 | 
|  |   1356 |           finally show ?thesis .
 | 
|  |   1357 |         next
 | 
|  |   1358 |           assume inf: "infinite (\<Union>i. F i)"
 | 
|  |   1359 |           { fix i
 | 
|  |   1360 |             have "\<exists>N. i \<le> (\<Sum>i<N. card (F i))"
 | 
|  |   1361 |             proof (induct i)
 | 
|  |   1362 |               case (Suc j)
 | 
|  |   1363 |               from Suc obtain N where N: "j \<le> (\<Sum>i<N. card (F i))" by auto
 | 
|  |   1364 |               have "infinite ({i. F i \<noteq> {}} - {..< N})"
 | 
|  |   1365 |                 using inf by (auto simp: fin_eq)
 | 
|  |   1366 |               then have "{i. F i \<noteq> {}} - {..< N} \<noteq> {}"
 | 
|  |   1367 |                 by (metis finite.emptyI)
 | 
|  |   1368 |               then obtain i where i: "F i \<noteq> {}" "N \<le> i"
 | 
|  |   1369 |                 by (auto simp: not_less[symmetric])
 | 
|  |   1370 | 
 | 
|  |   1371 |               note N
 | 
|  |   1372 |               also have "(\<Sum>i<N. card (F i)) \<le> (\<Sum>i<i. card (F i))"
 | 
|  |   1373 |                 by (rule setsum_mono2) (auto simp: i)
 | 
|  |   1374 |               also have "\<dots> < (\<Sum>i<i. card (F i)) + card (F i)"
 | 
|  |   1375 |                 using finite_F `F i \<noteq> {}` by (simp add: card_gt_0_iff)
 | 
|  |   1376 |               finally have "j < (\<Sum>i<Suc i. card (F i))"
 | 
|  |   1377 |                 by simp
 | 
|  |   1378 |               then show ?case unfolding Suc_le_eq by blast
 | 
|  |   1379 |             qed simp }
 | 
|  |   1380 |           with finite_F inf show ?thesis
 | 
|  |   1381 |             by (auto simp del: real_of_nat_setsum intro!: SUP_PInfty
 | 
|  |   1382 |                      simp add: suminf_ereal_eq_SUPR real_of_nat_setsum[symmetric])
 | 
|  |   1383 |         qed
 | 
|  |   1384 |       next
 | 
|  |   1385 |         assume "\<not> (\<forall>i. finite (F i))"
 | 
|  |   1386 |         then obtain j where j: "infinite (F j)" by auto
 | 
|  |   1387 |         then have "infinite (\<Union>i. F i)"
 | 
|  |   1388 |           using finite_subset[of "F j" "\<Union>i. F i"] by auto
 | 
|  |   1389 |         moreover have "\<And>i. 0 \<le> ?M (F i)" by auto
 | 
|  |   1390 |         ultimately show ?thesis
 | 
|  |   1391 |           using suminf_PInfty[of "\<lambda>i. ?M (F i)" j] j by auto
 | 
|  |   1392 |       qed
 | 
|  |   1393 |   qed
 | 
|  |   1394 |   show "X \<in> Pow A" using `X \<subseteq> A` by simp
 | 
|  |   1395 | qed
 | 
|  |   1396 | 
 | 
|  |   1397 | lemma emeasure_count_space_finite[simp]:
 | 
|  |   1398 |   "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = ereal (card X)"
 | 
|  |   1399 |   using emeasure_count_space[of X A] by simp
 | 
|  |   1400 | 
 | 
|  |   1401 | lemma emeasure_count_space_infinite[simp]:
 | 
|  |   1402 |   "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
 | 
|  |   1403 |   using emeasure_count_space[of X A] by simp
 | 
|  |   1404 | 
 | 
|  |   1405 | lemma emeasure_count_space_eq_0:
 | 
|  |   1406 |   "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
 | 
|  |   1407 | proof cases
 | 
|  |   1408 |   assume X: "X \<subseteq> A"
 | 
|  |   1409 |   then show ?thesis
 | 
|  |   1410 |   proof (intro iffI impI)
 | 
|  |   1411 |     assume "emeasure (count_space A) X = 0"
 | 
|  |   1412 |     with X show "X = {}"
 | 
|  |   1413 |       by (subst (asm) emeasure_count_space) (auto split: split_if_asm)
 | 
|  |   1414 |   qed simp
 | 
|  |   1415 | qed (simp add: emeasure_notin_sets)
 | 
|  |   1416 | 
 | 
|  |   1417 | lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
 | 
|  |   1418 |   unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
 | 
|  |   1419 | 
 | 
|  |   1420 | lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
 | 
|  |   1421 |   unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
 | 
|  |   1422 | 
 | 
|  |   1423 | lemma sigma_finite_measure_count_space:
 | 
|  |   1424 |   fixes A :: "'a::countable set"
 | 
|  |   1425 |   shows "sigma_finite_measure (count_space A)"
 | 
|  |   1426 | proof
 | 
|  |   1427 |   show "\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (count_space A) \<and> (\<Union>i. F i) = space (count_space A) \<and>
 | 
|  |   1428 |      (\<forall>i. emeasure (count_space A) (F i) \<noteq> \<infinity>)"
 | 
|  |   1429 |      using surj_from_nat by (intro exI[of _ "\<lambda>i. {from_nat i} \<inter> A"]) (auto simp del: surj_from_nat)
 | 
|  |   1430 | qed
 | 
|  |   1431 | 
 | 
|  |   1432 | lemma finite_measure_count_space:
 | 
|  |   1433 |   assumes [simp]: "finite A"
 | 
|  |   1434 |   shows "finite_measure (count_space A)"
 | 
|  |   1435 |   by rule simp
 | 
|  |   1436 | 
 | 
|  |   1437 | lemma sigma_finite_measure_count_space_finite:
 | 
|  |   1438 |   assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
 | 
|  |   1439 | proof -
 | 
|  |   1440 |   interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
 | 
|  |   1441 |   show "sigma_finite_measure (count_space A)" ..
 | 
|  |   1442 | qed
 | 
|  |   1443 | 
 | 
|  |   1444 | end
 | 
|  |   1445 | 
 |