# HG changeset patch # User hoelzl # Date 1443454141 -7200 # Node ID 542a4d9bac96dc0d721a3a3784e79fdd2c3b09ad # Parent f49644098959b0f2e6410234c4b4a1b38b99ccce Caratheodory: cleanup and modernisation diff -r f49644098959 -r 542a4d9bac96 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Fri Sep 25 23:01:34 2015 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Mon Sep 28 17:29:01 2015 +0200 @@ -41,49 +41,43 @@ ultimately show ?thesis unfolding g_def using pos by (auto intro!: SUP_eq simp: setsum.cartesian_product reindex SUP_upper2 - setsum_nonneg suminf_ereal_eq_SUP SUP_pair + suminf_ereal_eq_SUP SUP_pair SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg) qed subsection {* Characterizations of Measures *} -definition subadditive where "subadditive M f \ - (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" - -definition countably_subadditive where "countably_subadditive M f \ - (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ - (f (\i. A i) \ (\i. f (A i))))" +definition subadditive where + "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" -definition outer_measure_space where "outer_measure_space M f \ - positive M f \ increasing M f \ countably_subadditive M f" +definition countably_subadditive where + "countably_subadditive M f \ + (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" -definition measure_set where "measure_set M f X = {r. - \A. range A \ M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) = r}" +definition outer_measure_space where + "outer_measure_space M f \ positive M f \ increasing M f \ countably_subadditive M f" -lemma subadditiveD: - "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" +lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) subsubsection {* Lambda Systems *} -definition lambda_system where "lambda_system \ M f = {l \ M. - \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" +definition lambda_system where + "lambda_system \ M f = {l \ M. \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" lemma (in algebra) lambda_system_eq: - shows "lambda_system \ M f = {l \ M. \x \ M. f (x \ l) + f (x - l) = f x}" + "lambda_system \ M f = {l \ M. \x \ M. f (x \ l) + f (x - l) = f x}" proof - - have [simp]: "!!l x. l \ M \ x \ M \ (\ - l) \ x = x - l" + have [simp]: "\l x. l \ M \ x \ M \ (\ - l) \ x = x - l" by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) show ?thesis by (auto simp add: lambda_system_def) (metis Int_commute)+ qed -lemma (in algebra) lambda_system_empty: - "positive M f \ {} \ lambda_system \ M f" +lemma (in algebra) lambda_system_empty: "positive M f \ {} \ lambda_system \ M f" by (auto simp add: positive_def lambda_system_eq) -lemma lambda_system_sets: - "x \ lambda_system \ M f \ x \ M" +lemma lambda_system_sets: "x \ lambda_system \ M f \ x \ M" by (simp add: lambda_system_def) lemma (in algebra) lambda_system_Compl: @@ -201,12 +195,10 @@ by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed -lemma lambda_system_increasing: - "increasing M f \ increasing (lambda_system \ M f) f" +lemma lambda_system_increasing: "increasing M f \ increasing (lambda_system \ M f) f" by (simp add: increasing_def lambda_system_def) -lemma lambda_system_positive: - "positive M f \ positive (lambda_system \ M f) f" +lemma lambda_system_positive: "positive M f \ positive (lambda_system \ M f) f" by (simp add: positive_def lambda_system_def) lemma (in algebra) lambda_system_strong_sum: @@ -258,66 +250,56 @@ have *: "\i. 0 \ f (A i)" using pos A'' unfolding positive_def by auto have dis: "\N. disjoint_family_on A {..i. f (A i)) \ f (\i. A i)" - using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] - using A'' + using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A'' by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] countable_UN) qed - { - fix a - assume a [iff]: "a \ M" - have "f (a \ (\i. A i)) + f (a - (\i. A i)) = f a" - proof - - show ?thesis - proof (rule antisym) - have "range (\i. a \ A i) \ M" using A'' - by blast - moreover - have "disjoint_family (\i. a \ A i)" using disj - by (auto simp add: disjoint_family_on_def) - moreover - have "a \ (\i. A i) \ M" - by (metis Int U_in a) - ultimately - have "f (a \ (\i. A i)) \ (\i. f (a \ A i))" - using csa[unfolded countably_subadditive_def, rule_format, of "(\i. a \ A i)"] - by (simp add: o_def) - hence "f (a \ (\i. A i)) + f (a - (\i. A i)) \ - (\i. f (a \ A i)) + f (a - (\i. A i))" - by (rule add_right_mono) - moreover - have "(\i. f (a \ A i)) + f (a - (\i. A i)) \ f a" - proof (intro suminf_bound_add allI) - fix n - have UNION_in: "(\i\{0.. M" - by (metis A'' UNION_in_sets) - have le_fa: "f (UNION {0.. a) \ f a" using A'' - by (blast intro: increasingD [OF inc] A'' UNION_in_sets) - have ls: "(\i\{0.. lambda_system \ M f" - using ls.UNION_in_sets by (simp add: A) - hence eq_fa: "f a = f (a \ (\i\{0..i\{0..i. A i)) \ f (a - (\i\{0..i A i)) + f (a - (\i. A i)) \ f a" - by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) - next - have "\i. a \ A i \ M" using A'' by auto - then show "\i. 0 \ f (a \ A i)" using pos[unfolded positive_def] by auto - have "\i. a - (\i. A i) \ M" using A'' by auto - then have "\i. 0 \ f (a - (\i. A i))" using pos[unfolded positive_def] by auto - then show "f (a - (\i. A i)) \ -\" by auto - qed - ultimately show "f (a \ (\i. A i)) + f (a - (\i. A i)) \ f a" - by (rule order_trans) - next - have "f a \ f (a \ (\i. A i) \ (a - (\i. A i)))" - by (blast intro: increasingD [OF inc] U_in) - also have "... \ f (a \ (\i. A i)) + f (a - (\i. A i))" - by (blast intro: subadditiveD [OF sa] U_in) - finally show "f a \ f (a \ (\i. A i)) + f (a - (\i. A i))" . - qed - qed - } + have "f (a \ (\i. A i)) + f (a - (\i. A i)) = f a" + if a [iff]: "a \ M" for a + proof (rule antisym) + have "range (\i. a \ A i) \ M" using A'' + by blast + moreover + have "disjoint_family (\i. a \ A i)" using disj + by (auto simp add: disjoint_family_on_def) + moreover + have "a \ (\i. A i) \ M" + by (metis Int U_in a) + ultimately + have "f (a \ (\i. A i)) \ (\i. f (a \ A i))" + using csa[unfolded countably_subadditive_def, rule_format, of "(\i. a \ A i)"] + by (simp add: o_def) + hence "f (a \ (\i. A i)) + f (a - (\i. A i)) \ (\i. f (a \ A i)) + f (a - (\i. A i))" + by (rule add_right_mono) + also have "\ \ f a" + proof (intro suminf_bound_add allI) + fix n + have UNION_in: "(\i\{0.. M" + by (metis A'' UNION_in_sets) + have le_fa: "f (UNION {0.. a) \ f a" using A'' + by (blast intro: increasingD [OF inc] A'' UNION_in_sets) + have ls: "(\i\{0.. lambda_system \ M f" + using ls.UNION_in_sets by (simp add: A) + hence eq_fa: "f a = f (a \ (\i\{0..i\{0..i. A i)) \ f (a - (\i\{0..i A i)) + f (a - (\i. A i)) \ f a" + by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) + next + have "\i. a \ A i \ M" using A'' by auto + then show "\i. 0 \ f (a \ A i)" using pos[unfolded positive_def] by auto + have "\i. a - (\i. A i) \ M" using A'' by auto + then have "\i. 0 \ f (a - (\i. A i))" using pos[unfolded positive_def] by auto + then show "f (a - (\i. A i)) \ -\" by auto + qed + finally show "f (a \ (\i. A i)) + f (a - (\i. A i)) \ f a" . + next + have "f a \ f (a \ (\i. A i) \ (a - (\i. A i)))" + by (blast intro: increasingD [OF inc] U_in) + also have "... \ f (a \ (\i. A i)) + f (a - (\i. A i))" + by (blast intro: subadditiveD [OF sa] U_in) + finally show "f a \ f (a \ (\i. A i)) + f (a - (\i. A i))" . + qed thus ?thesis by (simp add: lambda_system_eq sums_iff U_eq U_in) qed @@ -345,300 +327,186 @@ using pos by (simp add: measure_space_def) qed -lemma inf_measure_nonempty: - assumes f: "positive M f" and b: "b \ M" and a: "a \ b" "{} \ M" - shows "f b \ measure_set M f a" -proof - - let ?A = "\i::nat. (if i = 0 then b else {})" - have "(\i. f (?A i)) = (\i<1::nat. f (?A i))" - by (rule suminf_finite) (simp_all add: f[unfolded positive_def]) - also have "... = f b" - by simp - finally show ?thesis using assms - by (auto intro!: exI [of _ ?A] - simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) -qed +definition outer_measure :: "'a set set \ ('a set \ ereal) \ 'a set \ ereal" where + "outer_measure M f X = + (INF A:{A. range A \ M \ disjoint_family A \ X \ (\i. A i)}. \i. f (A i))" -lemma (in ring_of_sets) inf_measure_agrees: - assumes posf: "positive M f" and ca: "countably_additive M f" - and s: "s \ M" - shows "Inf (measure_set M f s) = f s" -proof (intro Inf_eqI) - fix z - assume z: "z \ measure_set M f s" - from this obtain A where - A: "range A \ M" and disj: "disjoint_family A" - and "s \ (\x. A x)" and si: "(\i. f (A i)) = z" - by (auto simp add: measure_set_def comp_def) - hence seq: "s = (\i. A i \ s)" by blast +lemma (in ring_of_sets) outer_measure_agrees: + assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \ M" + shows "outer_measure M f s = f s" + unfolding outer_measure_def +proof (safe intro!: antisym INF_greatest) + fix A :: "nat \ 'a set" assume A: "range A \ M" and dA: "disjoint_family A" and sA: "s \ (\x. A x)" have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) - have sums: "(\i. f (A i \ s)) = f (\i. A i \ s)" - proof (rule ca[unfolded countably_additive_def, rule_format]) - show "range (\n. A n \ s) \ M" using A s - by blast - show "disjoint_family (\n. A n \ s)" using disj - by (auto simp add: disjoint_family_on_def) - show "(\i. A i \ s) \ M" using A s - by (metis UN_extend_simps(4) s seq) - qed - hence "f s = (\i. f (A i \ s))" - using seq [symmetric] by (simp add: sums_iff) + have "f s = f (\i. A i \ s)" + using sA by (auto simp: Int_absorb1) + also have "\ = (\i. f (A i \ s))" + using sA dA A s + by (intro ca[unfolded countably_additive_def, rule_format, symmetric]) + (auto simp: Int_absorb1 disjoint_family_on_def) also have "... \ (\i. f (A i))" - proof (rule suminf_le_pos) - fix n show "f (A n \ s) \ f (A n)" using A s - by (force intro: increasingD [OF inc]) - fix N have "A N \ s \ M" using A s by auto - then show "0 \ f (A N \ s)" using posf unfolding positive_def by auto - qed - also have "... = z" by (rule si) - finally show "f s \ z" . -qed (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl]) - -lemma measure_set_pos: - assumes posf: "positive M f" "r \ measure_set M f X" - shows "0 \ r" -proof - - obtain A where "range A \ M" and r: "r = (\i. f (A i))" - using `r \ measure_set M f X` unfolding measure_set_def by auto - then show "0 \ r" using posf unfolding r positive_def - by (intro suminf_0_le) auto -qed - -lemma inf_measure_pos: - assumes posf: "positive M f" - shows "0 \ Inf (measure_set M f X)" -proof (rule complete_lattice_class.Inf_greatest) - fix r assume "r \ measure_set M f X" with posf show "0 \ r" - by (rule measure_set_pos) + using A s by (intro suminf_le_pos increasingD[OF inc] positiveD2[OF posf]) auto + finally show "f s \ (\i. f (A i))" . +next + have "(\i. f (if i = 0 then s else {})) \ f s" + using positiveD1[OF posf] by (subst suminf_finite[of "{0}"]) auto + with s show "(INF A:{A. range A \ M \ disjoint_family A \ s \ UNION UNIV A}. \i. f (A i)) \ f s" + by (intro INF_lower2[of "\i. if i = 0 then s else {}"]) + (auto simp: disjoint_family_on_def) qed -lemma inf_measure_empty: +lemma outer_measure_nonneg: "positive M f \ 0 \ outer_measure M f X" + by (auto intro!: INF_greatest suminf_0_le intro: positiveD2 simp: outer_measure_def) + +lemma outer_measure_empty: assumes posf: "positive M f" and "{} \ M" - shows "Inf (measure_set M f {}) = 0" + shows "outer_measure M f {} = 0" proof (rule antisym) - show "Inf (measure_set M f {}) \ 0" - by (metis complete_lattice_class.Inf_lower `{} \ M` - inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) -qed (rule inf_measure_pos[OF posf]) + show "outer_measure M f {} \ 0" + using assms by (auto intro!: INF_lower2[of "\_. {}"] simp: outer_measure_def disjoint_family_on_def positive_def) +qed (intro outer_measure_nonneg posf) + +lemma (in ring_of_sets) positive_outer_measure: + assumes "positive M f" shows "positive (Pow \) (outer_measure M f)" + unfolding positive_def by (auto simp: assms outer_measure_empty outer_measure_nonneg) + +lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \) (outer_measure M f)" + by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower) -lemma (in ring_of_sets) inf_measure_positive: - assumes p: "positive M f" and "{} \ M" - shows "positive M (\x. Inf (measure_set M f x))" -proof (unfold positive_def, intro conjI ballI) - show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto - fix A assume "A \ M" -qed (rule inf_measure_pos[OF p]) +lemma (in ring_of_sets) outer_measure_le: + assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \ M" and X: "X \ (\i. A i)" + shows "outer_measure M f X \ (\i. f (A i))" + unfolding outer_measure_def +proof (safe intro!: INF_lower2[of "disjointed A"] del: subsetI) + show dA: "range (disjointed A) \ M" + by (auto intro!: A range_disjointed_sets) + have "\n. f (disjointed A n) \ f (A n)" + by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A) + moreover have "\i. 0 \ f (disjointed A i)" + using pos dA unfolding positive_def by auto + ultimately show "(\i. f (disjointed A i)) \ (\i. f (A i))" + by (blast intro!: suminf_le_pos) +qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed) -lemma (in ring_of_sets) inf_measure_increasing: - assumes posf: "positive M f" - shows "increasing (Pow \) (\x. Inf (measure_set M f x))" -apply (clarsimp simp add: increasing_def) -apply (rule complete_lattice_class.Inf_greatest) -apply (rule complete_lattice_class.Inf_lower) -apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast) -done - -lemma (in ring_of_sets) inf_measure_le: - assumes posf: "positive M f" and inc: "increasing M f" - and x: "x \ {r . \A. range A \ M \ s \ (\i. A i) \ (\i. f (A i)) = r}" - shows "Inf (measure_set M f s) \ x" +lemma (in ring_of_sets) outer_measure_close: + assumes posf: "positive M f" and "0 < e" and "outer_measure M f X \ \" + shows "\A. range A \ M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) \ outer_measure M f X + e" proof - - obtain A where A: "range A \ M" and ss: "s \ (\i. A i)" - and xeq: "(\i. f (A i)) = x" - using x by auto - have dA: "range (disjointed A) \ M" - by (metis A range_disjointed_sets) - have "\n. f (disjointed A n) \ f (A n)" - by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def) - moreover have "\i. 0 \ f (disjointed A i)" - using posf dA unfolding positive_def by auto - ultimately have sda: "(\i. f (disjointed A i)) \ (\i. f (A i))" - by (blast intro!: suminf_le_pos) - hence ley: "(\i. f (disjointed A i)) \ x" - by (metis xeq) - hence y: "(\i. f (disjointed A i)) \ measure_set M f s" - apply (auto simp add: measure_set_def) - apply (rule_tac x="disjointed A" in exI) - apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def) - done + from `outer_measure M f X \ \` have fin: "\outer_measure M f X\ \ \" + using outer_measure_nonneg[OF posf, of X] by auto show ?thesis - by (blast intro: y order_trans [OF _ ley] posf complete_lattice_class.Inf_lower) -qed - -lemma (in ring_of_sets) inf_measure_close: - fixes e :: ereal - assumes posf: "positive M f" and e: "0 < e" and ss: "s \ (\)" and "Inf (measure_set M f s) \ \" - shows "\A. range A \ M \ disjoint_family A \ s \ (\i. A i) \ - (\i. f (A i)) \ Inf (measure_set M f s) + e" -proof - - from `Inf (measure_set M f s) \ \` have fin: "\Inf (measure_set M f s)\ \ \" - using inf_measure_pos[OF posf, of s] by auto - obtain l where "l \ measure_set M f s" "l \ Inf (measure_set M f s) + e" - using Inf_ereal_close[OF fin e] by auto - thus ?thesis - by (auto intro!: exI[of _ l] simp: measure_set_def comp_def) + using Inf_ereal_close[OF fin[unfolded outer_measure_def INF_def], OF \0 < e\] + unfolding INF_def[symmetric] outer_measure_def[symmetric] by (auto intro: less_imp_le) qed -lemma (in ring_of_sets) inf_measure_countably_subadditive: +lemma (in ring_of_sets) countably_subadditive_outer_measure: assumes posf: "positive M f" and inc: "increasing M f" - shows "countably_subadditive (Pow \) (\x. Inf (measure_set M f x))" + shows "countably_subadditive (Pow \) (outer_measure M f)" proof (simp add: countably_subadditive_def, safe) - fix A :: "nat \ 'a set" - let ?outer = "\B. Inf (measure_set M f B)" - assume A: "range A \ Pow (\)" - and disj: "disjoint_family A" - and sb: "(\i. A i) \ \" + fix A :: "nat \ _" assume A: "range A \ Pow (\)" and sb: "(\i. A i) \ \" + let ?O = "outer_measure M f" - { fix e :: ereal assume e: "0 < e" and "\i. ?outer (A i) \ \" - hence "\BB. \n. range (BB n) \ M \ disjoint_family (BB n) \ - A n \ (\i. BB n i) \ (\i. f (BB n i)) \ ?outer (A n) + e * (1/2)^(Suc n)" - apply (safe intro!: choice inf_measure_close [of f, OF posf]) - using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def) - then obtain BB - where BB: "\n. (range (BB n) \ M)" - and disjBB: "\n. disjoint_family (BB n)" - and sbBB: "\n. A n \ (\i. BB n i)" - and BBle: "\n. (\i. f (BB n i)) \ ?outer (A n) + e * (1/2)^(Suc n)" + { fix e :: ereal assume e: "0 < e" and "\i. ?O (A i) \ \" + hence "\B. \n. range (B n) \ M \ disjoint_family (B n) \ A n \ (\i. B n i) \ + (\i. f (B n i)) \ ?O (A n) + e * (1/2)^(Suc n)" + using e sb by (auto intro!: choice outer_measure_close [of f, OF posf] simp: ereal_zero_less_0_iff one_ereal_def) + then obtain B + where B: "\n. range (B n) \ M" + and sbB: "\n. A n \ (\i. B n i)" + and Ble: "\n. (\i. f (B n i)) \ ?O (A n) + e * (1/2)^(Suc n)" by auto blast - have sll: "(\n. \i. (f (BB n i))) \ (\n. ?outer (A n)) + e" - proof - - have sum_eq_1: "(\n. e*(1/2) ^ Suc n) = e" - using suminf_half_series_ereal e - by (simp add: ereal_zero_le_0_iff zero_le_divide_ereal suminf_cmult_ereal) - have "\n i. 0 \ f (BB n i)" using posf[unfolded positive_def] BB by auto - then have "\n. 0 \ (\i. f (BB n i))" by (rule suminf_0_le) - then have "(\n. \i. (f (BB n i))) \ (\n. ?outer (A n) + e*(1/2) ^ Suc n)" - by (rule suminf_le_pos[OF BBle]) - also have "... = (\n. ?outer (A n)) + e" - using sum_eq_1 inf_measure_pos[OF posf] e - by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff) - finally show ?thesis . - qed - def C \ "(split BB) o prod_decode" - from BB have "\i j. BB i j \ M" + + def C \ "split B o prod_decode" + from B have B_in_M: "\i j. B i j \ M" by (rule range_subsetD) - then have C: "\n. C n \ M" - by (simp add: C_def split_def) - have sbC: "(\i. A i) \ (\i. C i)" - proof (auto simp add: C_def) - fix x i - assume x: "x \ A i" - with sbBB [of i] obtain j where "x \ BB i j" - by blast - thus "\i. x \ split BB (prod_decode i)" - by (metis prod_encode_inverse prod.case) - qed - have "(f \ C) = (f \ (\(x, y). BB x y)) \ prod_decode" - by (rule ext) (auto simp add: C_def) - moreover have "suminf ... = (\n. \i. f (BB n i))" using BBle - using BB posf[unfolded positive_def] - by (force intro!: suminf_ereal_2dimen simp: o_def) - ultimately have Csums: "(\i. f (C i)) = (\n. \i. f (BB n i))" by (simp add: o_def) - have "?outer (\i. A i) \ (\n. \i. f (BB n i))" - apply (rule inf_measure_le [OF posf(1) inc], auto) - apply (rule_tac x="C" in exI) - apply (auto simp add: C sbC Csums) - done - also have "... \ (\n. ?outer (A n)) + e" using sll - by blast - finally have "?outer (\i. A i) \ (\n. ?outer (A n)) + e" . } - note for_finite_Inf = this + then have C: "range C \ M" + by (auto simp add: C_def split_def) + have A_C: "(\i. A i) \ (\i. C i)" + using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse) - show "?outer (\i. A i) \ (\n. ?outer (A n))" + have "?O (\i. A i) \ ?O (\i. C i)" + using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space) + also have "\ \ (\i. f (C i))" + using C by (intro outer_measure_le[OF posf inc]) auto + also have "\ = (\n. \i. f (B n i))" + using B_in_M unfolding C_def comp_def by (intro suminf_ereal_2dimen positiveD2[OF posf]) auto + also have "\ \ (\n. ?O (A n) + e*(1/2) ^ Suc n)" + using B_in_M by (intro suminf_le_pos[OF Ble] suminf_0_le posf[THEN positiveD2]) auto + also have "... = (\n. ?O (A n)) + (\n. e*(1/2) ^ Suc n)" + using e by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff outer_measure_nonneg posf) + also have "(\n. e*(1/2) ^ Suc n) = e" + using suminf_half_series_ereal e by (simp add: ereal_zero_le_0_iff suminf_cmult_ereal) + finally have "?O (\i. A i) \ (\n. ?O (A n)) + e" . } + note * = this + + show "?O (\i. A i) \ (\n. ?O (A n))" proof cases - assume "\i. ?outer (A i) \ \" - with for_finite_Inf show ?thesis + assume "\i. ?O (A i) \ \" with * show ?thesis by (intro ereal_le_epsilon) auto - next - assume "\ (\i. ?outer (A i) \ \)" - then have "\i. ?outer (A i) = \" - by auto - then have "(\n. ?outer (A n)) = \" - using suminf_PInfty[OF inf_measure_pos, OF posf] - by metis - then show ?thesis by simp - qed + qed (metis suminf_PInfty[OF outer_measure_nonneg, OF posf] ereal_less_eq(1)) qed -lemma (in ring_of_sets) inf_measure_outer: - "\ positive M f ; increasing M f \ \ - outer_measure_space (Pow \) (\x. Inf (measure_set M f x))" - using inf_measure_pos[of M f] - by (simp add: outer_measure_space_def inf_measure_empty - inf_measure_increasing inf_measure_countably_subadditive positive_def) +lemma (in ring_of_sets) outer_measure_space_outer_measure: + "positive M f \ increasing M f \ outer_measure_space (Pow \) (outer_measure M f)" + by (simp add: outer_measure_space_def + positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure) lemma (in ring_of_sets) algebra_subset_lambda_system: assumes posf: "positive M f" and inc: "increasing M f" and add: "additive M f" - shows "M \ lambda_system \ (Pow \) (\x. Inf (measure_set M f x))" + shows "M \ lambda_system \ (Pow \) (outer_measure M f)" proof (auto dest: sets_into_space simp add: algebra.lambda_system_eq [OF algebra_Pow]) - fix x s - assume x: "x \ M" - and s: "s \ \" - have [simp]: "!!x. x \ M \ s \ (\ - x) = s-x" using s + fix x s assume x: "x \ M" and s: "s \ \" + have [simp]: "\x. x \ M \ s \ (\ - x) = s - x" using s by blast - have "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) - \ Inf (measure_set M f s)" - proof cases - assume "Inf (measure_set M f s) = \" then show ?thesis by simp - next - assume fin: "Inf (measure_set M f s) \ \" - then have "measure_set M f s \ {}" - by (auto simp: top_ereal_def) - show ?thesis - proof (rule complete_lattice_class.Inf_greatest) - fix r assume "r \ measure_set M f s" - then obtain A where A: "disjoint_family A" "range A \ M" "s \ (\i. A i)" - and r: "r = (\i. f (A i))" unfolding measure_set_def by auto - have "Inf (measure_set M f (s \ x)) \ (\i. f (A i \ x))" - unfolding measure_set_def - proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\i. A i \ x"]) - from A(1) show "disjoint_family (\i. A i \ x)" - by (rule disjoint_family_on_bisimulation) auto - qed (insert x A, auto) - moreover - have "Inf (measure_set M f (s - x)) \ (\i. f (A i - x))" - unfolding measure_set_def - proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\i. A i - x"]) - from A(1) show "disjoint_family (\i. A i - x)" - by (rule disjoint_family_on_bisimulation) auto - qed (insert x A, auto) - ultimately have "Inf (measure_set M f (s \ x)) + Inf (measure_set M f (s - x)) \ - (\i. f (A i \ x)) + (\i. f (A i - x))" by (rule add_mono) - also have "\ = (\i. f (A i \ x) + f (A i - x))" - using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def) - also have "\ = (\i. f (A i))" - using A x - by (subst add[THEN additiveD, symmetric]) - (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f]) - finally show "Inf (measure_set M f (s \ x)) + Inf (measure_set M f (s - x)) \ r" - using r by simp - qed + have "outer_measure M f (s \ x) + outer_measure M f (s - x) \ outer_measure M f s" + unfolding outer_measure_def[of M f s] + proof (safe intro!: INF_greatest) + fix A :: "nat \ 'a set" assume A: "disjoint_family A" "range A \ M" "s \ (\i. A i)" + have "outer_measure M f (s \ x) \ (\i. f (A i \ x))" + unfolding outer_measure_def + proof (safe intro!: INF_lower2[of "\i. A i \ x"]) + from A(1) show "disjoint_family (\i. A i \ x)" + by (rule disjoint_family_on_bisimulation) auto + qed (insert x A, auto) + moreover + have "outer_measure M f (s - x) \ (\i. f (A i - x))" + unfolding outer_measure_def + proof (safe intro!: INF_lower2[of "\i. A i - x"]) + from A(1) show "disjoint_family (\i. A i - x)" + by (rule disjoint_family_on_bisimulation) auto + qed (insert x A, auto) + ultimately have "outer_measure M f (s \ x) + outer_measure M f (s - x) \ + (\i. f (A i \ x)) + (\i. f (A i - x))" by (rule add_mono) + also have "\ = (\i. f (A i \ x) + f (A i - x))" + using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def) + also have "\ = (\i. f (A i))" + using A x + by (subst add[THEN additiveD, symmetric]) + (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f]) + finally show "outer_measure M f (s \ x) + outer_measure M f (s - x) \ (\i. f (A i))" . qed moreover - have "Inf (measure_set M f s) - \ Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x))" + have "outer_measure M f s \ outer_measure M f (s \ x) + outer_measure M f (s - x)" proof - - have "Inf (measure_set M f s) = Inf (measure_set M f ((s\x) \ (s-x)))" + have "outer_measure M f s = outer_measure M f ((s \ x) \ (s - x))" by (metis Un_Diff_Int Un_commute) - also have "... \ Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x))" + also have "... \ outer_measure M f (s \ x) + outer_measure M f (s - x)" apply (rule subadditiveD) apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow]) - apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf]) - apply (rule inf_measure_countably_subadditive) + apply (simp add: positive_def outer_measure_empty[OF posf] outer_measure_nonneg[OF posf]) + apply (rule countably_subadditive_outer_measure) using s by (auto intro!: posf inc) finally show ?thesis . qed ultimately - show "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) - = Inf (measure_set M f s)" + show "outer_measure M f (s \ x) + outer_measure M f (s - x) = outer_measure M f s" by (rule order_antisym) qed -lemma measure_down: - "measure_space \ N \ \ sigma_algebra \ M \ M \ N \ measure_space \ M \" +lemma measure_down: "measure_space \ N \ \ sigma_algebra \ M \ M \ N \ measure_space \ M \" by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq) subsection {* Caratheodory's theorem *} @@ -649,11 +517,11 @@ proof - have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) - let ?infm = "(\x. Inf (measure_set M f x))" - def ls \ "lambda_system \ (Pow \) ?infm" - have mls: "measure_space \ ls ?infm" + let ?O = "outer_measure M f" + def ls \ "lambda_system \ (Pow \) ?O" + have mls: "measure_space \ ls ?O" using sigma_algebra.caratheodory_lemma - [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]] + [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]] by (simp add: ls_def) hence sls: "sigma_algebra \ ls" by (simp add: measure_space_def) @@ -663,11 +531,11 @@ hence sgs_sb: "sigma_sets (\) (M) \ ls" using sigma_algebra.sigma_sets_subset [OF sls, of "M"] by simp - have "measure_space \ (sigma_sets \ M) ?infm" + have "measure_space \ (sigma_sets \ M) ?O" by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets) (simp_all add: sgs_sb space_closed) - thus ?thesis using inf_measure_agrees [OF posf ca] - by (intro exI[of _ ?infm]) auto + thus ?thesis using outer_measure_agrees [OF posf ca] + by (intro exI[of _ ?O]) auto qed lemma (in ring_of_sets) caratheodory_empty_continuous: @@ -1069,5 +937,4 @@ then show ?thesis by simp qed - end