# HG changeset patch # User hoelzl # Date 1296646485 -3600 # Node ID 3e39b0e730d6f3579551c8d9fd34eeae358b2591 # Parent f9ff311992b6dae7f9cee08114fe47e2c7104c21 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale; changed syntax for simple_function, simple_integral, positive_integral, integral and RN_deriv. introduced binder variants for simple_integral, positive_integral and integral. diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Caratheodory.thy Wed Feb 02 12:34:45 2011 +0100 @@ -8,77 +8,72 @@ subsection {* Measure Spaces *} -definition "positive f \ f {} = (0::pextreal)" -- "Positive is enforced by the type" +record 'a measure_space = "'a algebra" + + measure :: "'a set \ pextreal" -definition - additive where - "additive M f \ - (\x \ sets M. \y \ sets M. x \ y = {} - \ f (x \ y) = f x + f y)" +definition positive where "positive M f \ f {} = (0::pextreal)" + -- "Positive is enforced by the type" -definition - countably_additive where - "countably_additive M f \ - (\A. range A \ sets M \ - disjoint_family A \ - (\i. A i) \ sets M \ - (\\<^isub>\ n. f (A n)) = f (\i. A i))" +definition additive where "additive M f \ + (\x \ sets M. \y \ sets M. x \ y = {} \ f (x \ y) = f x + f y)" -definition - increasing where - "increasing M f \ (\x \ sets M. \y \ sets M. x \ y \ f x \ f y)" +definition countably_additive where "countably_additive M f \ + (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ + (\\<^isub>\ n. f (A n)) = f (\i. A i))" -definition - subadditive where - "subadditive M f \ - (\x \ sets M. \y \ sets M. x \ y = {} - \ f (x \ y) \ f x + f y)" +definition increasing where "increasing M f \ + (\x \ sets M. \y \ sets M. x \ y \ f x \ f y)" -definition - countably_subadditive where - "countably_subadditive M f \ - (\A. range A \ sets M \ - disjoint_family A \ - (\i. A i) \ sets M \ - f (\i. A i) \ psuminf (\n. f (A n)))" +definition subadditive where "subadditive M f \ + (\x \ sets M. \y \ sets M. x \ y = {} \ + f (x \ y) \ f x + f y)" -definition - lambda_system where - "lambda_system M f = - {l. l \ sets M & (\x \ sets M. f (l \ x) + f ((space M - l) \ x) = f x)}" +definition countably_subadditive where "countably_subadditive M f \ + (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ + f (\i. A i) \ (\\<^isub>\ n. f (A n)))" + +definition lambda_system where "lambda_system M f = {l \ sets M. + \x \ sets M. f (l \ x) + f ((space M - l) \ x) = f x}" -definition - outer_measure_space where - "outer_measure_space M f \ - positive f \ increasing M f \ countably_subadditive M f" +definition outer_measure_space where "outer_measure_space M f \ + positive M f \ increasing M f \ countably_subadditive M f" + +definition measure_set where "measure_set M f X = {r. + \A. range A \ sets M \ disjoint_family A \ X \ (\i. A i) \ (\\<^isub>\ i. f (A i)) = r}" -definition - measure_set where - "measure_set M f X = - {r . \A. range A \ sets M \ disjoint_family A \ X \ (\i. A i) \ (\\<^isub>\ i. f (A i)) = r}" +locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" + + assumes empty_measure [simp]: "measure M {} = 0" + and ca: "countably_additive M (measure M)" -locale measure_space = sigma_algebra + - fixes \ :: "'a set \ pextreal" - assumes empty_measure [simp]: "\ {} = 0" - and ca: "countably_additive M \" +abbreviation (in measure_space) "\ \ measure M" lemma increasingD: - "increasing M f \ x \ y \ x\sets M \ y\sets M \ f x \ f y" + "increasing M f \ x \ y \ x\sets M \ y\sets M \ f x \ f y" by (auto simp add: increasing_def) lemma subadditiveD: - "subadditive M f \ x \ y = {} \ x\sets M \ y\sets M - \ f (x \ y) \ f x + f y" + "subadditive M f \ x \ y = {} \ x \ sets M \ y \ sets M + \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) lemma additiveD: - "additive M f \ x \ y = {} \ x\sets M \ y\sets M - \ f (x \ y) = f x + f y" + "additive M f \ x \ y = {} \ x \ sets M \ y \ sets M + \ f (x \ y) = f x + f y" by (auto simp add: additive_def) lemma countably_additiveD: "countably_additive M f \ range A \ sets M \ disjoint_family A - \ (\i. A i) \ sets M \ (\\<^isub>\ n. f (A n)) = f (\i. A i)" + \ (\i. A i) \ sets M \ (\\<^isub>\ n. f (A n)) = f (\i. A i)" + by (simp add: countably_additive_def) + +lemma countably_subadditiveD: + "countably_subadditive M f \ range A \ sets M \ disjoint_family A \ + (\i. A i) \ sets M \ f (\i. A i) \ psuminf (f o A)" + by (auto simp add: countably_subadditive_def o_def) + +lemma countably_additiveI: + "(\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M + \ (\\<^isub>\ n. f (A n)) = f (\i. A i)) \ countably_additive M f" by (simp add: countably_additive_def) section "Extend binary sets" @@ -109,7 +104,8 @@ by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f]) lemma suminf_binaryset_eq: - "f {} = 0 \ suminf (\n. f (binaryset A B n)) = f A + f B" + fixes f :: "'a set \ real" + shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) lemma binaryset_psuminf: @@ -130,8 +126,8 @@ subsection {* Lambda Systems *} lemma (in algebra) lambda_system_eq: - "lambda_system M f = - {l. l \ sets M & (\x \ sets M. f (x \ l) + f (x - l) = f x)}" + shows "lambda_system M f = {l \ sets M. + \x \ sets M. f (x \ l) + f (x - l) = f x}" proof - have [simp]: "!!l x. l \ sets M \ x \ sets M \ (space M - l) \ x = x - l" by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) @@ -140,60 +136,59 @@ qed lemma (in algebra) lambda_system_empty: - "positive f \ {} \ lambda_system M f" - by (auto simp add: positive_def lambda_system_eq) + "positive M f \ {} \ lambda_system M f" + by (auto simp add: positive_def lambda_system_eq algebra_def) lemma lambda_system_sets: - "x \ lambda_system M f \ x \ sets M" - by (simp add: lambda_system_def) + "x \ lambda_system M f \ x \ sets M" + by (simp add: lambda_system_def) lemma (in algebra) lambda_system_Compl: fixes f:: "'a set \ pextreal" assumes x: "x \ lambda_system M f" shows "space M - x \ lambda_system M f" - proof - - have "x \ space M" - by (metis sets_into_space lambda_system_sets x) - hence "space M - (space M - x) = x" - by (metis double_diff equalityE) - with x show ?thesis - by (force simp add: lambda_system_def ac_simps) - qed +proof - + have "x \ space M" + by (metis sets_into_space lambda_system_sets x) + hence "space M - (space M - x) = x" + by (metis double_diff equalityE) + with x show ?thesis + by (force simp add: lambda_system_def ac_simps) +qed lemma (in algebra) lambda_system_Int: fixes f:: "'a set \ pextreal" assumes xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" shows "x \ y \ lambda_system M f" - proof - - from xl yl show ?thesis - proof (auto simp add: positive_def lambda_system_eq Int) - fix u - assume x: "x \ sets M" and y: "y \ sets M" and u: "u \ sets M" - and fx: "\z\sets M. f (z \ x) + f (z - x) = f z" - and fy: "\z\sets M. f (z \ y) + f (z - y) = f z" - have "u - x \ y \ sets M" - by (metis Diff Diff_Int Un u x y) - moreover - have "(u - (x \ y)) \ y = u \ y - x" by blast - moreover - have "u - x \ y - y = u - y" by blast - ultimately - have ey: "f (u - x \ y) = f (u \ y - x) + f (u - y)" using fy - by force - have "f (u \ (x \ y)) + f (u - x \ y) - = (f (u \ (x \ y)) + f (u \ y - x)) + f (u - y)" - by (simp add: ey ac_simps) - also have "... = (f ((u \ y) \ x) + f (u \ y - x)) + f (u - y)" - by (simp add: Int_ac) - also have "... = f (u \ y) + f (u - y)" - using fx [THEN bspec, of "u \ y"] Int y u - by force - also have "... = f u" - by (metis fy u) - finally show "f (u \ (x \ y)) + f (u - x \ y) = f u" . - qed +proof - + from xl yl show ?thesis + proof (auto simp add: positive_def lambda_system_eq Int) + fix u + assume x: "x \ sets M" and y: "y \ sets M" and u: "u \ sets M" + and fx: "\z\sets M. f (z \ x) + f (z - x) = f z" + and fy: "\z\sets M. f (z \ y) + f (z - y) = f z" + have "u - x \ y \ sets M" + by (metis Diff Diff_Int Un u x y) + moreover + have "(u - (x \ y)) \ y = u \ y - x" by blast + moreover + have "u - x \ y - y = u - y" by blast + ultimately + have ey: "f (u - x \ y) = f (u \ y - x) + f (u - y)" using fy + by force + have "f (u \ (x \ y)) + f (u - x \ y) + = (f (u \ (x \ y)) + f (u \ y - x)) + f (u - y)" + by (simp add: ey ac_simps) + also have "... = (f ((u \ y) \ x) + f (u \ y - x)) + f (u - y)" + by (simp add: Int_ac) + also have "... = f (u \ y) + f (u - y)" + using fx [THEN bspec, of "u \ y"] Int y u + by force + also have "... = f u" + by (metis fy u) + finally show "f (u \ (x \ y)) + f (u - x \ y) = f u" . qed - +qed lemma (in algebra) lambda_system_Un: fixes f:: "'a set \ pextreal" @@ -210,7 +205,7 @@ qed lemma (in algebra) lambda_system_algebra: - "positive f \ algebra (M (|sets := lambda_system M f|))" + "positive M f \ algebra (M\sets := lambda_system M f\)" apply (auto simp add: algebra_def) apply (metis lambda_system_sets set_mp sets_into_space) apply (metis lambda_system_empty) @@ -222,32 +217,31 @@ assumes z: "z \ sets M" and disj: "x \ y = {}" and xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" shows "f (z \ (x \ y)) = f (z \ x) + f (z \ y)" - proof - - have "z \ x = (z \ (x \ y)) \ x" using disj by blast - moreover - have "z \ y = (z \ (x \ y)) - x" using disj by blast - moreover - have "(z \ (x \ y)) \ sets M" - by (metis Int Un lambda_system_sets xl yl z) - ultimately show ?thesis using xl yl - by (simp add: lambda_system_eq) - qed +proof - + have "z \ x = (z \ (x \ y)) \ x" using disj by blast + moreover + have "z \ y = (z \ (x \ y)) - x" using disj by blast + moreover + have "(z \ (x \ y)) \ sets M" + by (metis Int Un lambda_system_sets xl yl z) + ultimately show ?thesis using xl yl + by (simp add: lambda_system_eq) +qed lemma (in algebra) lambda_system_additive: "additive (M (|sets := lambda_system M f|)) f" - proof (auto simp add: additive_def) - fix x and y - assume disj: "x \ y = {}" - and xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" - hence "x \ sets M" "y \ sets M" by (blast intro: lambda_system_sets)+ - thus "f (x \ y) = f x + f y" - using lambda_system_strong_additive [OF top disj xl yl] - by (simp add: Un) - qed - +proof (auto simp add: additive_def) + fix x and y + assume disj: "x \ y = {}" + and xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" + hence "x \ sets M" "y \ sets M" by (blast intro: lambda_system_sets)+ + thus "f (x \ y) = f x + f y" + using lambda_system_strong_additive [OF top disj xl yl] + by (simp add: Un) +qed lemma (in algebra) countably_subadditive_subadditive: - assumes f: "positive f" and cs: "countably_subadditive M f" + assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y @@ -267,7 +261,7 @@ lemma (in algebra) additive_sum: fixes A:: "nat \ 'a set" - assumes f: "positive f" and ad: "additive M f" + assumes f: "positive M f" and ad: "additive M f" and A: "range A \ sets M" and disj: "disjoint_family A" shows "setsum (f \ A) {0..i\{0.. range A \ sets M \ disjoint_family A \ - (\i. A i) \ sets M \ f (\i. A i) \ psuminf (f o A)" - by (auto simp add: countably_subadditive_def o_def) - lemma (in algebra) increasing_additive_bound: fixes A:: "nat \ 'a set" and f :: "'a set \ pextreal" - assumes f: "positive f" and ad: "additive M f" + assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ sets M" and disj: "disjoint_family A" @@ -311,12 +299,16 @@ qed lemma lambda_system_increasing: - "increasing M f \ increasing (M (|sets := lambda_system M f|)) f" + "increasing M f \ increasing (M (|sets := lambda_system M f|)) f" by (simp add: increasing_def lambda_system_def) +lemma lambda_system_positive: + "positive M f \ positive (M (|sets := lambda_system M f|)) f" + by (simp add: positive_def lambda_system_def) + lemma (in algebra) lambda_system_strong_sum: fixes A:: "nat \ 'a set" and f :: "'a set \ pextreal" - assumes f: "positive f" and a: "a \ sets M" + assumes f: "positive M f" and a: "a \ sets M" and A: "range A \ lambda_system M f" and disj: "disjoint_family A" shows "(\i = 0..A i)) = f (a \ (\i\{0.. lambda_system M f" and disj: "disjoint_family A" shows "(\i. A i) \ lambda_system M f \ psuminf (f \ A) = f (\i. A i)" proof - - have pos: "positive f" and inc: "increasing M f" + have pos: "positive M f" and inc: "increasing M f" and csa: "countably_subadditive M f" by (metis oms outer_measure_space_def)+ have sa: "subadditive M f" @@ -357,15 +348,15 @@ have U_in: "(\i. A i) \ sets M" by (metis A'' countable_UN) have U_eq: "f (\i. A i) = psuminf (f o A)" - proof (rule antisym) - show "f (\i. A i) \ psuminf (f \ A)" - by (rule countably_subadditiveD [OF csa A'' disj U_in]) - show "psuminf (f \ A) \ f (\i. A i)" - by (rule psuminf_bound, unfold atLeast0LessThan[symmetric]) - (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right - lambda_system_additive subset_Un_eq increasingD [OF inc] - A' A'' UNION_in_sets U_in) - qed + proof (rule antisym) + show "f (\i. A i) \ psuminf (f \ A)" + by (rule countably_subadditiveD [OF csa A'' disj U_in]) + show "psuminf (f \ A) \ f (\i. A i)" + by (rule psuminf_bound, unfold atLeast0LessThan[symmetric]) + (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right + lambda_system_positive lambda_system_additive + subset_Un_eq increasingD [OF inc] A' A'' UNION_in_sets U_in) + qed { fix a assume a [iff]: "a \ sets M" @@ -424,19 +415,20 @@ lemma (in sigma_algebra) caratheodory_lemma: assumes oms: "outer_measure_space M f" - shows "measure_space (|space = space M, sets = lambda_system M f|) f" + shows "measure_space \ space = space M, sets = lambda_system M f, measure = f \" + (is "measure_space ?M") proof - - have pos: "positive f" + have pos: "positive M f" by (metis oms outer_measure_space_def) - have alg: "algebra (|space = space M, sets = lambda_system M f|)" + have alg: "algebra ?M" using lambda_system_algebra [of f, OF pos] by (simp add: algebra_def) then moreover - have "sigma_algebra (|space = space M, sets = lambda_system M f|)" + have "sigma_algebra ?M" using lambda_system_caratheodory [OF oms] by (simp add: sigma_algebra_disjoint_iff) moreover - have "measure_space_axioms (|space = space M, sets = lambda_system M f|) f" + have "measure_space_axioms ?M" using pos lambda_system_caratheodory [OF oms] by (simp add: measure_space_axioms_def positive_def lambda_system_sets countably_additive_def o_def) @@ -446,7 +438,7 @@ qed lemma (in algebra) additive_increasing: - assumes posf: "positive f" and addf: "additive M f" + assumes posf: "positive M f" and addf: "additive M f" shows "increasing M f" proof (auto simp add: increasing_def) fix x y @@ -460,7 +452,7 @@ qed lemma (in algebra) countably_additive_additive: - assumes posf: "positive f" and ca: "countably_additive M f" + assumes posf: "positive M f" and ca: "countably_additive M f" shows "additive M f" proof (auto simp add: additive_def) fix x y @@ -480,7 +472,7 @@ qed lemma inf_measure_nonempty: - assumes f: "positive f" and b: "b \ sets M" and a: "a \ b" "{} \ sets M" + assumes f: "positive M f" and b: "b \ sets M" and a: "a \ b" "{} \ sets M" shows "f b \ measure_set M f a" proof - have "psuminf (f \ (\i. {})(0 := b)) = setsum (f \ (\i. {})(0 := b)) {..<1::nat}" @@ -494,7 +486,7 @@ qed lemma (in algebra) inf_measure_agrees: - assumes posf: "positive f" and ca: "countably_additive M f" + assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \ sets M" shows "Inf (measure_set M f s) = f s" unfolding Inf_pextreal_def @@ -530,25 +522,25 @@ fix y assume y: "\u \ measure_set M f s. y \ u" thus "y \ f s" - by (blast intro: inf_measure_nonempty [of f, OF posf s subset_refl]) + by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl]) qed -lemma (in algebra) inf_measure_empty: - assumes posf: "positive f" "{} \ sets M" +lemma inf_measure_empty: + assumes posf: "positive M f" "{} \ sets M" shows "Inf (measure_set M f {}) = 0" proof (rule antisym) show "Inf (measure_set M f {}) \ 0" - by (metis complete_lattice_class.Inf_lower `{} \ sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) + by (metis complete_lattice_class.Inf_lower `{} \ sets M` + inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) qed simp lemma (in algebra) inf_measure_positive: - "positive f \ - positive (\x. Inf (measure_set M f x))" - by (simp add: positive_def inf_measure_empty) + "positive M f \ positive M (\x. Inf (measure_set M f x))" + by (simp add: positive_def inf_measure_empty) lemma (in algebra) inf_measure_increasing: - assumes posf: "positive f" - shows "increasing (| space = space M, sets = Pow (space M) |) + assumes posf: "positive M f" + shows "increasing \ space = space M, sets = Pow (space M) \ (\x. Inf (measure_set M f x))" apply (auto simp add: increasing_def) apply (rule complete_lattice_class.Inf_greatest) @@ -558,7 +550,7 @@ lemma (in algebra) inf_measure_le: - assumes posf: "positive f" and inc: "increasing M f" + assumes posf: "positive M f" and inc: "increasing M f" and x: "x \ {r . \A. range A \ sets M \ s \ (\i. A i) \ psuminf (f \ A) = r}" shows "Inf (measure_set M f s) \ x" proof - @@ -584,7 +576,7 @@ qed lemma (in algebra) inf_measure_close: - assumes posf: "positive f" and e: "0 < e" and ss: "s \ (space M)" + assumes posf: "positive M f" and e: "0 < e" and ss: "s \ (space M)" shows "\A. range A \ sets M \ disjoint_family A \ s \ (\i. A i) \ psuminf (f \ A) \ Inf (measure_set M f s) + e" proof (cases "Inf (measure_set M f s) = \") @@ -596,7 +588,7 @@ next case True have "measure_set M f s \ {}" - by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets]) + by (metis emptyE ss inf_measure_nonempty [of _ f, OF posf top _ empty_sets]) then obtain l where "l \ measure_set M f s" by auto moreover from True have "l \ Inf (measure_set M f s) + e" by simp ultimately show ?thesis @@ -604,7 +596,7 @@ qed lemma (in algebra) inf_measure_countably_subadditive: - assumes posf: "positive f" and inc: "increasing M f" + assumes posf: "positive M f" and inc: "increasing M f" shows "countably_subadditive (| space = space M, sets = Pow (space M) |) (\x. Inf (measure_set M f x))" unfolding countably_subadditive_def o_def @@ -666,8 +658,8 @@ qed lemma (in algebra) inf_measure_outer: - "\ positive f ; increasing M f \ - \ outer_measure_space (| space = space M, sets = Pow (space M) |) + "\ positive M f ; increasing M f \ + \ outer_measure_space \ space = space M, sets = Pow (space M) \ (\x. Inf (measure_set M f x))" by (simp add: outer_measure_space_def inf_measure_empty inf_measure_increasing inf_measure_countably_subadditive positive_def) @@ -675,7 +667,7 @@ (*MOVE UP*) lemma (in algebra) algebra_subset_lambda_system: - assumes posf: "positive f" and inc: "increasing M f" + assumes posf: "positive M f" and inc: "increasing M f" and add: "additive M f" shows "sets M \ lambda_system (| space = space M, sets = Pow (space M) |) (\x. Inf (measure_set M f x))" @@ -739,10 +731,10 @@ by (metis Un_Diff_Int Un_commute) also have "... \ Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x))" apply (rule subadditiveD) - apply (iprover intro: algebra.countably_subadditive_subadditive algebra_Pow - inf_measure_positive inf_measure_countably_subadditive posf inc) - apply (auto simp add: subsetD [OF s]) - done + apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow]) + apply (simp add: positive_def inf_measure_empty[OF posf]) + apply (rule inf_measure_countably_subadditive) + using s by (auto intro!: posf inc) finally show ?thesis . qed ultimately @@ -752,37 +744,38 @@ qed lemma measure_down: - "measure_space N \ \ sigma_algebra M \ sets M \ sets N \ - (\ = \) \ measure_space M \" + "measure_space N \ sigma_algebra M \ sets M \ sets N \ measure N = measure M \ measure_space M" by (simp add: measure_space_def measure_space_axioms_def positive_def countably_additive_def) blast theorem (in algebra) caratheodory: - assumes posf: "positive f" and ca: "countably_additive M f" - shows "\\ :: 'a set \ pextreal. (\s \ sets M. \ s = f s) \ measure_space (sigma M) \" - 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 (|space = space M, sets = Pow (space M)|) ?infm" - have mls: "measure_space \space = space M, sets = ls\ ?infm" - using sigma_algebra.caratheodory_lemma - [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]] - by (simp add: ls_def) - hence sls: "sigma_algebra (|space = space M, sets = ls|)" - by (simp add: measure_space_def) - have "sets M \ ls" - by (simp add: ls_def) - (metis ca posf inc countably_additive_additive algebra_subset_lambda_system) - hence sgs_sb: "sigma_sets (space M) (sets M) \ ls" - using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"] - by simp - have "measure_space (sigma M) ?infm" - unfolding sigma_def - 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 (auto intro!: exI[of _ ?infm]) - qed + assumes posf: "positive M f" and ca: "countably_additive M f" + shows "\\ :: 'a set \ pextreal. (\s \ sets M. \ s = f s) \ + measure_space \ space = space M, sets = sets (sigma M), measure = \ \" +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 (|space = space M, sets = Pow (space M)|) ?infm" + have mls: "measure_space \space = space M, sets = ls, measure = ?infm\" + using sigma_algebra.caratheodory_lemma + [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]] + by (simp add: ls_def) + hence sls: "sigma_algebra (|space = space M, sets = ls, measure = ?infm|)" + by (simp add: measure_space_def) + have "sets M \ ls" + by (simp add: ls_def) + (metis ca posf inc countably_additive_additive algebra_subset_lambda_system) + hence sgs_sb: "sigma_sets (space M) (sets M) \ ls" + using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"] + by simp + have "measure_space \ space = space M, sets = sets (sigma M), measure = ?infm \" + unfolding sigma_def + 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 +qed end diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Complete_Measure.thy Wed Feb 02 12:34:45 2011 +0100 @@ -7,9 +7,24 @@ locale completeable_measure_space = measure_space -definition (in completeable_measure_space) completion :: "'a algebra" where +definition (in completeable_measure_space) + "split_completion A p = (\N'. A = fst p \ snd p \ fst p \ snd p = {} \ + fst p \ sets M \ snd p \ N' \ N' \ null_sets)" + +definition (in completeable_measure_space) + "main_part A = fst (Eps (split_completion A))" + +definition (in completeable_measure_space) + "null_part A = snd (Eps (split_completion A))" + +abbreviation (in completeable_measure_space) "\' A \ \ (main_part A)" + +definition (in completeable_measure_space) completion :: "('a, 'b) measure_space_scheme" where "completion = \ space = space M, - sets = { S \ N |S N N'. S \ sets M \ N' \ null_sets \ N \ N' } \" + sets = { S \ N |S N N'. S \ sets M \ N' \ null_sets \ N \ N' }, + measure = \', + \ = more M \" + lemma (in completeable_measure_space) space_completion[simp]: "space completion = space M" unfolding completion_def by simp @@ -58,16 +73,6 @@ auto qed auto -definition (in completeable_measure_space) - "split_completion A p = (\N'. A = fst p \ snd p \ fst p \ snd p = {} \ - fst p \ sets M \ snd p \ N' \ N' \ null_sets)" - -definition (in completeable_measure_space) - "main_part A = fst (Eps (split_completion A))" - -definition (in completeable_measure_space) - "null_part A = snd (Eps (split_completion A))" - lemma (in completeable_measure_space) split_completion: assumes "A \ sets completion" shows "split_completion A (main_part A, null_part A)" @@ -108,17 +113,15 @@ show "\ (null_part S) = 0" by auto qed -definition (in completeable_measure_space) "\' A = \ (main_part A)" - lemma (in completeable_measure_space) \'_set[simp]: assumes "S \ sets M" shows "\' S = \ S" proof - have S: "S \ sets completion" using assms by auto then have "\ S = \ (main_part S \ null_part S)" by simp - also have "\ = \ (main_part S)" + also have "\ = \' S" using S assms measure_additive[of "main_part S" "null_part S"] by (auto simp: measure_additive) - finally show ?thesis unfolding \'_def by simp + finally show ?thesis by simp qed lemma (in completeable_measure_space) sets_completionI_sub: @@ -154,7 +157,7 @@ unfolding * .. also have "\ = \ (\i. main_part (S i))" using null_set S by (intro measure_Un_null_set) auto - finally show ?thesis unfolding \'_def . + finally show ?thesis . qed lemma (in completeable_measure_space) \_main_part_Un: @@ -168,30 +171,35 @@ unfolding range_binary_eq Un_range_binary UN by auto qed -sublocale completeable_measure_space \ completion!: measure_space completion \' -proof - show "\' {} = 0" by auto -next - show "countably_additive completion \'" - proof (unfold countably_additive_def, intro allI conjI impI) - fix A :: "nat \ 'a set" assume A: "range A \ sets completion" "disjoint_family A" - have "disjoint_family (\i. main_part (A i))" - proof (intro disjoint_family_on_bisimulation[OF A(2)]) - fix n m assume "A n \ A m = {}" - then have "(main_part (A n) \ null_part (A n)) \ (main_part (A m) \ null_part (A m)) = {}" - using A by (subst (1 2) main_part_null_part_Un) auto - then show "main_part (A n) \ main_part (A m) = {}" by auto +sublocale completeable_measure_space \ completion!: measure_space completion + where "measure completion = \'" +proof - + show "measure_space completion" + proof + show "measure completion {} = 0" by (auto simp: completion_def) + next + show "countably_additive completion (measure completion)" + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume A: "range A \ sets completion" "disjoint_family A" + have "disjoint_family (\i. main_part (A i))" + proof (intro disjoint_family_on_bisimulation[OF A(2)]) + fix n m assume "A n \ A m = {}" + then have "(main_part (A n) \ null_part (A n)) \ (main_part (A m) \ null_part (A m)) = {}" + using A by (subst (1 2) main_part_null_part_Un) auto + then show "main_part (A n) \ main_part (A m) = {}" by auto + qed + then have "(\\<^isub>\n. measure completion (A n)) = \ (\i. main_part (A i))" + unfolding completion_def using A by (auto intro!: measure_countably_additive) + then show "(\\<^isub>\n. measure completion (A n)) = measure completion (UNION UNIV A)" + by (simp add: completion_def \_main_part_UN[OF A(1)]) qed - then have "(\\<^isub>\n. \' (A n)) = \ (\i. main_part (A i))" - unfolding \'_def using A by (intro measure_countably_additive) auto - then show "(\\<^isub>\n. \' (A n)) = \' (UNION UNIV A)" - unfolding \_main_part_UN[OF A(1)] . qed + show "measure completion = \'" unfolding completion_def by simp qed lemma (in completeable_measure_space) completion_ex_simple_function: - assumes f: "completion.simple_function f" - shows "\f'. simple_function f' \ (AE x. f x = f' x)" + assumes f: "simple_function completion f" + shows "\f'. simple_function M f' \ (AE x. f x = f' x)" proof - let "?F x" = "f -` {x} \ space M" have F: "\x. ?F x \ sets completion" and fin: "finite (f`space M)" @@ -248,11 +256,11 @@ shows "\g'\borel_measurable M. (AE x. g x = g' x)" proof - from g[THEN completion.borel_measurable_implies_simple_function_sequence] - obtain f where "\i. completion.simple_function (f i)" "f \ g" by auto - then have "\i. \f'. simple_function f' \ (AE x. f i x = f' x)" + obtain f where "\i. simple_function completion (f i)" "f \ g" by auto + then have "\i. \f'. simple_function M f' \ (AE x. f i x = f' x)" using completion_ex_simple_function by auto from this[THEN choice] obtain f' where - sf: "\i. simple_function (f' i)" and + sf: "\i. simple_function M (f' i)" and AE: "\i. AE x. f i x = f' i x" by auto show ?thesis proof (intro bexI) diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Information.thy Wed Feb 02 12:34:45 2011 +0100 @@ -165,43 +165,45 @@ Kullback$-$Leibler distance. *} definition - "KL_divergence b M \ \ = - measure_space.integral M \ (\x. log b (real (sigma_finite_measure.RN_deriv M \ \ x)))" + "KL_divergence b M \ = \x. log b (real (RN_deriv M \ x)) \M\measure := \\" lemma (in sigma_finite_measure) KL_divergence_cong: - assumes "measure_space M \" - and cong: "\A. A \ sets M \ \' A = \ A" "\A. A \ sets M \ \' A = \ A" - shows "KL_divergence b M \' \' = KL_divergence b M \ \" + assumes "measure_space (M\measure := \\)" (is "measure_space ?\") + assumes [simp]: "sets N = sets M" "space N = space M" + "\A. A \ sets M \ measure N A = \ A" + "\A. A \ sets M \ \ A = \' A" + shows "KL_divergence b M \ = KL_divergence b N \'" proof - - interpret \: measure_space M \ by fact - show ?thesis - unfolding KL_divergence_def - using RN_deriv_cong[OF cong, of "\A. A"] - by (simp add: cong \.integral_cong_measure[OF cong(2)]) + interpret \: measure_space ?\ by fact + have "KL_divergence b M \ = \x. log b (real (RN_deriv N \' x)) \?\" + by (simp cong: RN_deriv_cong \.integral_cong add: KL_divergence_def) + also have "\ = KL_divergence b N \'" + by (auto intro!: \.integral_cong_measure[symmetric] simp: KL_divergence_def) + finally show ?thesis . qed lemma (in finite_measure_space) KL_divergence_eq_finite: - assumes v: "finite_measure_space M \" + assumes v: "finite_measure_space (M\measure := \\)" assumes ac: "absolutely_continuous \" - shows "KL_divergence b M \ \ = (\x\space M. real (\ {x}) * log b (real (\ {x}) / real (\ {x})))" (is "_ = ?sum") + shows "KL_divergence b M \ = (\x\space M. real (\ {x}) * log b (real (\ {x}) / real (\ {x})))" (is "_ = ?sum") proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v]) - interpret v: finite_measure_space M \ by fact - have ms: "measure_space M \" by fact - show "(\x \ space M. log b (real (RN_deriv \ x)) * real (\ {x})) = ?sum" + interpret v: finite_measure_space "M\measure := \\" by fact + have ms: "measure_space (M\measure := \\)" by default + show "(\x \ space M. log b (real (RN_deriv M \ x)) * real (\ {x})) = ?sum" using RN_deriv_finite_measure[OF ms ac] by (auto intro!: setsum_cong simp: field_simps real_of_pextreal_mult[symmetric]) qed lemma (in finite_prob_space) KL_divergence_positive_finite: - assumes v: "finite_prob_space M \" + assumes v: "finite_prob_space (M\measure := \\)" assumes ac: "absolutely_continuous \" and "1 < b" - shows "0 \ KL_divergence b M \ \" + shows "0 \ KL_divergence b M \" proof - - interpret v: finite_prob_space M \ using v . - have ms: "finite_measure_space M \" by default + interpret v: finite_prob_space "M\measure := \\" by fact + have ms: "finite_measure_space (M\measure := \\)" by default - have "- (KL_divergence b M \ \) \ log b (\x\space M. real (\ {x}))" + have "- (KL_divergence b M \) \ log b (\x\space M. real (\ {x}))" proof (subst KL_divergence_eq_finite[OF ms ac], safe intro!: log_setsum_divide not_empty) show "finite (space M)" using finite_space by simp show "1 < b" by fact @@ -215,16 +217,15 @@ using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto thus "0 < prob {x}" using finite_measure[of "{x}"] x by auto } qed auto - thus "0 \ KL_divergence b M \ \" using finite_sum_over_space_eq_1 by simp + thus "0 \ KL_divergence b M \" using finite_sum_over_space_eq_1 by simp qed subsection {* Mutual Information *} definition (in prob_space) "mutual_information b S T X Y = - KL_divergence b (sigma (pair_algebra S T)) - (joint_distribution X Y) - (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y))" + KL_divergence b (S\measure := distribution X\ \\<^isub>M T\measure := distribution Y\) + (joint_distribution X Y)" definition (in prob_space) "entropy b s X = mutual_information b s s X X" @@ -232,32 +233,49 @@ abbreviation (in information_space) mutual_information_Pow ("\'(_ ; _')") where "\(X ; Y) \ mutual_information b - \ space = X`space M, sets = Pow (X`space M) \ - \ space = Y`space M, sets = Pow (Y`space M) \ X Y" + \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ X Y" + +lemma algebra_measure_update[simp]: + "algebra (M'\measure := m\) \ algebra M'" + unfolding algebra_def by simp + +lemma sigma_algebra_measure_update[simp]: + "sigma_algebra (M'\measure := m\) \ sigma_algebra M'" + unfolding sigma_algebra_def sigma_algebra_axioms_def by simp + +lemma finite_sigma_algebra_measure_update[simp]: + "finite_sigma_algebra (M'\measure := m\) \ finite_sigma_algebra M'" + unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp lemma (in prob_space) finite_variables_absolutely_continuous: assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" - shows "measure_space.absolutely_continuous (sigma (pair_algebra S T)) - (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y)) (joint_distribution X Y)" + shows "measure_space.absolutely_continuous + (S\measure := distribution X\ \\<^isub>M T\measure := distribution Y\) + (joint_distribution X Y)" proof - - interpret X: finite_prob_space S "distribution X" using X by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space T "distribution Y" using Y by (rule distribution_finite_prob_space) - interpret XY: pair_finite_prob_space S "distribution X" T "distribution Y" by default - interpret P: finite_prob_space XY.P "joint_distribution X Y" - using assms by (intro joint_distribution_finite_prob_space) + interpret X: finite_prob_space "S\measure := distribution X\" + using X by (rule distribution_finite_prob_space) + interpret Y: finite_prob_space "T\measure := distribution Y\" + using Y by (rule distribution_finite_prob_space) + interpret XY: pair_finite_prob_space + "S\measure := distribution X\" "T\ measure := distribution Y\" by default + interpret P: finite_prob_space "XY.P\ measure := joint_distribution X Y\" + using assms by (auto intro!: joint_distribution_finite_prob_space) + note rv = assms[THEN finite_random_variableD] show "XY.absolutely_continuous (joint_distribution X Y)" proof (rule XY.absolutely_continuousI) - show "finite_measure_space XY.P (joint_distribution X Y)" by default - fix x assume "x \ space XY.P" and "XY.pair_measure {x} = 0" + show "finite_measure_space (XY.P\ measure := joint_distribution X Y\)" by default + fix x assume "x \ space XY.P" and "XY.\ {x} = 0" then obtain a b where "(a, b) = x" and "a \ space S" "b \ space T" and distr: "distribution X {a} * distribution Y {b} = 0" - by (cases x) (auto simp: pair_algebra_def) - with assms[THEN finite_random_variableD] - joint_distribution_Times_le_fst[of S X T Y "{a}" "{b}"] - joint_distribution_Times_le_snd[of S X T Y "{a}" "{b}"] + by (cases x) (auto simp: space_pair_measure) + with X.sets_eq_Pow Y.sets_eq_Pow + joint_distribution_Times_le_fst[OF rv, of "{a}" "{b}"] + joint_distribution_Times_le_snd[OF rv, of "{a}" "{b}"] have "joint_distribution X Y {x} \ distribution Y {b}" "joint_distribution X Y {x} \ distribution X {a}" - by auto + by (auto simp del: X.sets_eq_Pow Y.sets_eq_Pow) with distr show "joint_distribution X Y {x} = 0" by auto qed qed @@ -274,19 +292,21 @@ and mutual_information_positive_generic: "0 \ mutual_information b MX MY X Y" (is ?positive) proof - - interpret X: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space MY "distribution Y" using MY by (rule distribution_finite_prob_space) - interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y" by default - interpret P: finite_prob_space XY.P "joint_distribution X Y" - using assms by (intro joint_distribution_finite_prob_space) + interpret X: finite_prob_space "MX\measure := distribution X\" + using MX by (rule distribution_finite_prob_space) + interpret Y: finite_prob_space "MY\measure := distribution Y\" + using MY by (rule distribution_finite_prob_space) + interpret XY: pair_finite_prob_space "MX\measure := distribution X\" "MY\measure := distribution Y\" by default + interpret P: finite_prob_space "XY.P\measure := joint_distribution X Y\" + using assms by (auto intro!: joint_distribution_finite_prob_space) - have P_ms: "finite_measure_space XY.P (joint_distribution X Y)" by default - have P_ps: "finite_prob_space XY.P (joint_distribution X Y)" by default + have P_ms: "finite_measure_space (XY.P\measure :=joint_distribution X Y\)" by default + have P_ps: "finite_prob_space (XY.P\measure := joint_distribution X Y\)" by default show ?sum unfolding Let_def mutual_information_def by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]]) - (auto simp add: pair_algebra_def setsum_cartesian_product' real_of_pextreal_mult[symmetric]) + (auto simp add: space_pair_measure setsum_cartesian_product' real_of_pextreal_mult[symmetric]) show ?positive using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1] @@ -301,12 +321,12 @@ by (auto simp add: ac_simps intro!: setsum_reindex_cong[OF swap_inj_on]) lemma (in information_space) mutual_information_commute_simple: - assumes X: "simple_function X" and Y: "simple_function Y" + assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X;Y) = \(Y;X)" by (intro X Y simple_function_imp_finite_random_variable mutual_information_commute) lemma (in information_space) mutual_information_eq: - assumes "simple_function X" "simple_function Y" + assumes "simple_function M X" "simple_function M Y" shows "\(X;Y) = (\ (x,y) \ X ` space M \ Y ` space M. real (distribution (\x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\x. (X x, Y x)) {(x,y)}) / (real (distribution X {x}) * real (distribution Y {y}))))" @@ -327,7 +347,7 @@ by (simp cong: distribution_cong image_cong) lemma (in information_space) mutual_information_positive: - assumes "simple_function X" "simple_function Y" + assumes "simple_function M X" "simple_function M Y" shows "0 \ \(X;Y)" using assms by (simp add: mutual_information_positive_generic) @@ -335,13 +355,14 @@ abbreviation (in information_space) entropy_Pow ("\'(_')") where - "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M) \ X" + "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ X" lemma (in information_space) entropy_generic_eq: assumes MX: "finite_random_variable MX X" shows "entropy b MX X = -(\ x \ space MX. real (distribution X {x}) * log b (real (distribution X {x})))" proof - - interpret MX: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space) + interpret MX: finite_prob_space "MX\measure := distribution X\" + using MX by (rule distribution_finite_prob_space) let "?X x" = "real (distribution X {x})" let "?XX x y" = "real (joint_distribution X X {(x, y)})" { fix x y @@ -353,25 +374,26 @@ show ?thesis unfolding entropy_def mutual_information_generic_eq[OF MX MX] unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX - by (auto simp: setsum_cases MX.finite_space) + using MX.finite_space by (auto simp: setsum_cases) qed lemma (in information_space) entropy_eq: - assumes "simple_function X" + assumes "simple_function M X" shows "\(X) = -(\ x \ X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))" using assms by (simp add: entropy_generic_eq) lemma (in information_space) entropy_positive: - "simple_function X \ 0 \ \(X)" + "simple_function M X \ 0 \ \(X)" unfolding entropy_def by (simp add: mutual_information_positive) lemma (in information_space) entropy_certainty_eq_0: - assumes "simple_function X" and "x \ X ` space M" and "distribution X {x} = 1" + assumes "simple_function M X" and "x \ X ` space M" and "distribution X {x} = 1" shows "\(X) = 0" proof - - interpret X: finite_prob_space "\ space = X ` space M, sets = Pow (X ` space M) \" "distribution X" - using simple_function_imp_finite_random_variable[OF `simple_function X`] - by (rule distribution_finite_prob_space) + let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" + note simple_function_imp_finite_random_variable[OF `simple_function M X`] + from distribution_finite_prob_space[OF this, of "\ measure = distribution X \"] + interpret X: finite_prob_space ?X by simp have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" using X.measure_compl[of "{x}"] assms by auto also have "\ = 0" using X.prob_space assms by auto @@ -383,38 +405,39 @@ hence fi: "\ y. y \ X ` space M \ real (distribution X {y}) = (if x = y then 1 else 0)" using assms by auto have y: "\y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp - show ?thesis unfolding entropy_eq[OF `simple_function X`] by (auto simp: y fi) + show ?thesis unfolding entropy_eq[OF `simple_function M X`] by (auto simp: y fi) qed lemma (in information_space) entropy_le_card_not_0: - assumes "simple_function X" + assumes "simple_function M X" shows "\(X) \ log b (real (card (X ` space M \ {x . distribution X {x} \ 0})))" proof - let "?d x" = "distribution X {x}" let "?p x" = "real (?d x)" have "\(X) = (\x\X`space M. ?p x * log b (1 / ?p x))" - by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function X`] setsum_negf[symmetric] log_simps not_less) + by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function M X`] setsum_negf[symmetric] log_simps not_less) also have "\ \ log b (\x\X`space M. ?p x * (1 / ?p x))" apply (rule log_setsum') - using not_empty b_gt_1 `simple_function X` sum_over_space_real_distribution + using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution by (auto simp: simple_function_def) also have "\ = log b (\x\X`space M. if ?d x \ 0 then 1 else 0)" - using distribution_finite[OF `simple_function X`[THEN simple_function_imp_random_variable], simplified] + using distribution_finite[OF `simple_function M X`[THEN simple_function_imp_random_variable], simplified] by (intro arg_cong[where f="\X. log b X"] setsum_cong) (auto simp: real_of_pextreal_eq_0) finally show ?thesis - using `simple_function X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def) + using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def) qed lemma (in information_space) entropy_uniform_max: - assumes "simple_function X" + assumes "simple_function M X" assumes "\x y. \ x \ X ` space M ; y \ X ` space M \ \ distribution X {x} = distribution X {y}" shows "\(X) = log b (real (card (X ` space M)))" proof - - interpret X: finite_prob_space "\ space = X ` space M, sets = Pow (X ` space M) \" "distribution X" - using simple_function_imp_finite_random_variable[OF `simple_function X`] - by (rule distribution_finite_prob_space) + let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" + note simple_function_imp_finite_random_variable[OF `simple_function M X`] + from distribution_finite_prob_space[OF this, of "\ measure = distribution X \"] + interpret X: finite_prob_space ?X by simp have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff - using `simple_function X` not_empty by (auto simp: simple_function_def) + using `simple_function M X` not_empty by (auto simp: simple_function_def) { fix x assume "x \ X ` space M" hence "real (distribution X {x}) = 1 / real (card (X ` space M))" proof (rule X.uniform_prob[simplified]) @@ -423,18 +446,18 @@ qed } thus ?thesis using not_empty X.finite_space b_gt_1 card_gt0 - by (simp add: entropy_eq[OF `simple_function X`] real_eq_of_nat[symmetric] log_simps) + by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps) qed lemma (in information_space) entropy_le_card: - assumes "simple_function X" + assumes "simple_function M X" shows "\(X) \ log b (real (card (X ` space M)))" proof cases assume "X ` space M \ {x. distribution X {x} \ 0} = {}" then have "\x. x\X`space M \ distribution X {x} = 0" by auto moreover have "0 < card (X`space M)" - using `simple_function X` not_empty + using `simple_function M X` not_empty by (auto simp: card_gt_0_iff simple_function_def) then have "log b 1 \ log b (real (card (X`space M)))" using b_gt_1 by (intro log_le) auto @@ -451,10 +474,10 @@ qed lemma (in information_space) entropy_commute: - assumes "simple_function X" "simple_function Y" + assumes "simple_function M X" "simple_function M Y" shows "\(\x. (X x, Y x)) = \(\x. (Y x, X x))" proof - - have sf: "simple_function (\x. (X x, Y x))" "simple_function (\x. (Y x, X x))" + have sf: "simple_function M (\x. (X x, Y x))" "simple_function M (\x. (Y x, X x))" using assms by (auto intro: simple_function_Pair) have *: "(\x. (Y x, X x))`space M = (\(a,b). (b,a))`(\x. (X x, Y x))`space M" by auto @@ -466,12 +489,12 @@ qed lemma (in information_space) entropy_eq_cartesian_product: - assumes "simple_function X" "simple_function Y" + assumes "simple_function M X" "simple_function M Y" shows "\(\x. (X x, Y x)) = -(\x\X`space M. \y\Y`space M. real (joint_distribution X Y {(x,y)}) * log b (real (joint_distribution X Y {(x,y)})))" proof - - have sf: "simple_function (\x. (X x, Y x))" + have sf: "simple_function M (\x. (X x, Y x))" using assms by (auto intro: simple_function_Pair) { fix x assume "x\(\x. (X x, Y x))`space M" then have "(\x. (X x, Y x)) -` {x} \ space M = {}" by auto @@ -485,19 +508,18 @@ subsection {* Conditional Mutual Information *} definition (in prob_space) - "conditional_mutual_information b M1 M2 M3 X Y Z \ - mutual_information b M1 (sigma (pair_algebra M2 M3)) X (\x. (Y x, Z x)) - - mutual_information b M1 M3 X Z" + "conditional_mutual_information b MX MY MZ X Y Z \ + mutual_information b MX (MY \\<^isub>M MZ) X (\x. (Y x, Z x)) - + mutual_information b MX MZ X Z" abbreviation (in information_space) conditional_mutual_information_Pow ("\'( _ ; _ | _ ')") where "\(X ; Y | Z) \ conditional_mutual_information b - \ space = X`space M, sets = Pow (X`space M) \ - \ space = Y`space M, sets = Pow (Y`space M) \ - \ space = Z`space M, sets = Pow (Z`space M) \ + \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ + \ space = Z`space M, sets = Pow (Z`space M), measure = distribution Z \ X Y Z" - lemma (in information_space) conditional_mutual_information_generic_eq: assumes MX: "finite_random_variable MX X" and MY: "finite_random_variable MY Y" @@ -519,7 +541,7 @@ note finite_var = MX MY MZ note random_var = finite_var[THEN finite_random_variableD] - note space_simps = space_pair_algebra space_sigma algebra.simps + note space_simps = space_pair_measure space_sigma algebra.simps note YZ = finite_random_variable_pairI[OF finite_var(2,3)] note XZ = finite_random_variable_pairI[OF finite_var(1,3)] @@ -574,12 +596,12 @@ unfolding conditional_mutual_information_def unfolding mutual_information_generic_eq[OF finite_var(1,3)] unfolding mutual_information_generic_eq[OF finite_var(1) YZ] - by (simp add: space_sigma space_pair_algebra setsum_cartesian_product') + by (simp add: space_sigma space_pair_measure setsum_cartesian_product') finally show ?thesis by simp qed lemma (in information_space) conditional_mutual_information_eq: - assumes "simple_function X" "simple_function Y" "simple_function Z" + assumes "simple_function M X" "simple_function M Y" "simple_function M Z" shows "\(X;Y|Z) = (\(x, y, z) \ X`space M \ Y`space M \ Z`space M. real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) * log b (real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) / @@ -588,11 +610,11 @@ by simp lemma (in information_space) conditional_mutual_information_eq_mutual_information: - assumes X: "simple_function X" and Y: "simple_function Y" + assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X ; Y) = \(X ; Y | (\x. ()))" proof - have [simp]: "(\x. ()) ` space M = {()}" using not_empty by auto - have C: "simple_function (\x. ())" by auto + have C: "simple_function M (\x. ())" by auto show ?thesis unfolding conditional_mutual_information_eq[OF X Y C] unfolding mutual_information_eq[OF X Y] @@ -608,12 +630,13 @@ lemma (in prob_space) setsum_distribution: assumes X: "finite_random_variable MX X" shows "(\a\space MX. distribution X {a}) = 1" using setsum_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV \" "\x. ()" "{()}"] - using sigma_algebra_Pow[of "UNIV::unit set"] by simp + using sigma_algebra_Pow[of "UNIV::unit set" "()"] by simp lemma (in prob_space) setsum_real_distribution: + fixes MX :: "('c, 'd) measure_space_scheme" assumes X: "finite_random_variable MX X" shows "(\a\space MX. real (distribution X {a})) = 1" - using setsum_real_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV \" "\x. ()" "{()}"] - using sigma_algebra_Pow[of "UNIV::unit set"] by simp + using setsum_real_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV, measure = undefined \" "\x. ()" "{()}"] + using sigma_algebra_Pow[of "UNIV::unit set" "\ measure = undefined \"] by simp lemma (in information_space) conditional_mutual_information_generic_positive: assumes "finite_random_variable MX X" and "finite_random_variable MY Y" and "finite_random_variable MZ Z" @@ -633,7 +656,7 @@ have split_beta: "\f. split f = (\x. f (fst x) (snd x))" by (simp add: fun_eq_iff) - note space_simps = space_pair_algebra space_sigma algebra.simps + note space_simps = space_pair_measure space_sigma algebra.simps note finite_var = assms note YZ = finite_random_variable_pairI[OF finite_var(2,3)] @@ -672,7 +695,7 @@ unfolding setsum_cartesian_product' unfolding setsum_commute[of _ "space MY"] unfolding setsum_commute[of _ "space MZ"] - by (simp_all add: space_pair_algebra + by (simp_all add: space_pair_measure setsum_real_joint_distribution_singleton[OF `finite_random_variable MX X` YZ] setsum_real_joint_distribution_singleton[OF `finite_random_variable MY Y` finite_var(3)] setsum_real_distribution[OF `finite_random_variable MZ Z`]) @@ -704,10 +727,9 @@ qed lemma (in information_space) conditional_mutual_information_positive: - assumes "simple_function X" and "simple_function Y" and "simple_function Z" + assumes "simple_function M X" and "simple_function M Y" and "simple_function M Z" shows "0 \ \(X;Y|Z)" - using conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]] - by simp + by (rule conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]]) subsection {* Conditional Entropy *} @@ -717,16 +739,17 @@ abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where "\(X | Y) \ conditional_entropy b - \ space = X`space M, sets = Pow (X`space M) \ - \ space = Y`space M, sets = Pow (Y`space M) \ X Y" + \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ X Y" lemma (in information_space) conditional_entropy_positive: - "simple_function X \ simple_function Y \ 0 \ \(X | Y)" + "simple_function M X \ simple_function M Y \ 0 \ \(X | Y)" unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive) lemma (in measure_space) empty_measureI: "A = {} \ \ A = 0" by simp lemma (in information_space) conditional_entropy_generic_eq: + fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme" assumes MX: "finite_random_variable MX X" assumes MZ: "finite_random_variable MZ Z" shows "conditional_entropy b MX MZ X Z = @@ -743,7 +766,7 @@ { fix x z have "?XXZ x x z = ?XZ x z" unfolding distribution_def by (auto intro!: arg_cong[where f=\]) } note this[simp] - { fix x x' :: 'b and z assume "x' \ x" + { fix x x' :: 'c and z assume "x' \ x" then have "?XXZ x x' z = 0" by (auto simp: distribution_def intro!: arg_cong[where f=\] empty_measureI) } note this[simp] @@ -762,7 +785,6 @@ finally have "(\x'\space MX. real (?XXZ x x' z) * ?f x x' z) = - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" . } note * = this - show ?thesis unfolding conditional_entropy_def unfolding conditional_mutual_information_generic_eq[OF MX MX MZ] @@ -772,7 +794,7 @@ qed lemma (in information_space) conditional_entropy_eq: - assumes "simple_function X" "simple_function Z" + assumes "simple_function M X" "simple_function M Z" shows "\(X | Z) = - (\(x, z)\X ` space M \ Z ` space M. real (joint_distribution X Z {(x, z)}) * @@ -781,7 +803,7 @@ by simp lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis: - assumes X: "simple_function X" and Y: "simple_function Y" + assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X | Y) = -(\y\Y`space M. real (distribution Y {y}) * (\x\X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) * @@ -794,7 +816,7 @@ intro!: setsum_cong) lemma (in information_space) conditional_entropy_eq_cartesian_product: - assumes "simple_function X" "simple_function Y" + assumes "simple_function M X" "simple_function M Y" shows "\(X | Y) = -(\x\X`space M. \y\Y`space M. real (joint_distribution X Y {(x,y)}) * log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))" @@ -804,7 +826,7 @@ subsection {* Equalities *} lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: - assumes X: "simple_function X" and Z: "simple_function Z" + assumes X: "simple_function M X" and Z: "simple_function M Z" shows "\(X ; Z) = \(X) - \(X | Z)" proof - let "?XZ x z" = "real (joint_distribution X Z {(x, z)})" @@ -828,7 +850,7 @@ qed lemma (in information_space) conditional_entropy_less_eq_entropy: - assumes X: "simple_function X" and Z: "simple_function Z" + assumes X: "simple_function M X" and Z: "simple_function M Z" shows "\(X | Z) \ \(X)" proof - have "\(X ; Z) = \(X) - \(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] . @@ -837,7 +859,7 @@ qed lemma (in information_space) entropy_chain_rule: - assumes X: "simple_function X" and Y: "simple_function Y" + assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(\x. (X x, Y x)) = \(X) + \(Y|X)" proof - let "?XY x y" = "real (joint_distribution X Y {(x, y)})" @@ -976,7 +998,7 @@ qed lemma (in information_space) entropy_partition: - assumes sf: "simple_function X" "simple_function P" + assumes sf: "simple_function M X" "simple_function M P" assumes svi: "subvimage (space M) X P" shows "\(X) = \(P) + \(X|P)" proof - @@ -1026,10 +1048,10 @@ qed corollary (in information_space) entropy_data_processing: - assumes X: "simple_function X" shows "\(f \ X) \ \(X)" + assumes X: "simple_function M X" shows "\(f \ X) \ \(X)" proof - note X - moreover have fX: "simple_function (f \ X)" using X by auto + moreover have fX: "simple_function M (f \ X)" using X by auto moreover have "subvimage (space M) X (f \ X)" by auto ultimately have "\(X) = \(f\X) + \(X|f\X)" by (rule entropy_partition) then show "\(f \ X) \ \(X)" @@ -1037,12 +1059,12 @@ qed corollary (in information_space) entropy_of_inj: - assumes X: "simple_function X" and inj: "inj_on f (X`space M)" + assumes X: "simple_function M X" and inj: "inj_on f (X`space M)" shows "\(f \ X) = \(X)" proof (rule antisym) show "\(f \ X) \ \(X)" using entropy_data_processing[OF X] . next - have sf: "simple_function (f \ X)" + have sf: "simple_function M (f \ X)" using X by auto have "\(X) = \(the_inv_into (X`space M) f \ (f \ X))" by (auto intro!: mutual_information_cong simp: entropy_def the_inv_into_f_f[OF inj]) diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 12:34:45 2011 +0100 @@ -7,6 +7,7 @@ begin lemma sums_If_finite: + fixes f :: "nat \ 'a::real_normed_vector" assumes finite: "finite {r. P r}" shows "(\r. if P r then f r else 0) sums (\r\{r. P r}. f r)" (is "?F sums _") proof cases @@ -24,7 +25,8 @@ qed lemma sums_single: - "(\r. if r = i then f r else 0) sums f i" + fixes f :: "nat \ 'a::real_normed_vector" + shows "(\r. if r = i then f r else 0) sums f i" using sums_If_finite[of "\r. r = i" f] by simp section "Simple function" @@ -37,12 +39,12 @@ *} -definition (in sigma_algebra) "simple_function g \ +definition "simple_function M g \ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" lemma (in sigma_algebra) simple_functionD: - assumes "simple_function g" + assumes "simple_function M g" shows "finite (g ` space M)" and "g -` X \ space M \ sets M" proof - show "finite (g ` space M)" @@ -55,7 +57,7 @@ lemma (in sigma_algebra) simple_function_indicator_representation: fixes f ::"'a \ pextreal" - assumes f: "simple_function f" and x: "x \ space M" + assumes f: "simple_function M f" and x: "x \ space M" shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" (is "?l = ?r") proof - @@ -69,7 +71,7 @@ qed lemma (in measure_space) simple_function_notspace: - "simple_function (\x. h x * indicator (- space M) x::pextreal)" (is "simple_function ?h") + "simple_function M (\x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h") proof - have "?h ` space M \ {0}" unfolding indicator_def by auto hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) @@ -79,7 +81,7 @@ lemma (in sigma_algebra) simple_function_cong: assumes "\t. t \ space M \ f t = g t" - shows "simple_function f \ simple_function g" + shows "simple_function M f \ simple_function M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" @@ -87,15 +89,21 @@ thus ?thesis unfolding simple_function_def using assms by simp qed +lemma (in sigma_algebra) simple_function_cong_algebra: + assumes "sets N = sets M" "space N = space M" + shows "simple_function M f \ simple_function N f" + unfolding simple_function_def assms .. + lemma (in sigma_algebra) borel_measurable_simple_function: - assumes "simple_function f" + assumes "simple_function M f" shows "f \ borel_measurable M" proof (rule borel_measurableI) fix S let ?I = "f ` (f -` S \ space M)" have *: "(\x\?I. f -` {x} \ space M) = f -` S \ space M" (is "?U = _") by auto have "finite ?I" - using assms unfolding simple_function_def by (auto intro: finite_subset) + using assms unfolding simple_function_def + using finite_subset[of "f ` (f -` S \ space M)" "f ` space M"] by auto hence "?U \ sets M" apply (rule finite_UN) using assms unfolding simple_function_def by auto @@ -105,17 +113,17 @@ lemma (in sigma_algebra) simple_function_borel_measurable: fixes f :: "'a \ 'x::t2_space" assumes "f \ borel_measurable M" and "finite (f ` space M)" - shows "simple_function f" + shows "simple_function M f" using assms unfolding simple_function_def by (auto intro: borel_measurable_vimage) lemma (in sigma_algebra) simple_function_const[intro, simp]: - "simple_function (\x. c)" + "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) lemma (in sigma_algebra) simple_function_compose[intro, simp]: - assumes "simple_function f" - shows "simple_function (g \ f)" + assumes "simple_function M f" + shows "simple_function M (g \ f)" unfolding simple_function_def proof safe show "finite ((g \ f) ` space M)" @@ -132,7 +140,7 @@ lemma (in sigma_algebra) simple_function_indicator[intro, simp]: assumes "A \ sets M" - shows "simple_function (indicator A)" + shows "simple_function M (indicator A)" proof - have "indicator A ` space M \ {0, 1}" (is "?S \ _") by (auto simp: indicator_def) @@ -143,9 +151,9 @@ qed lemma (in sigma_algebra) simple_function_Pair[intro, simp]: - assumes "simple_function f" - assumes "simple_function g" - shows "simple_function (\x. (f x, g x))" (is "simple_function ?p") + assumes "simple_function M f" + assumes "simple_function M g" + shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") unfolding simple_function_def proof safe show "finite (?p ` space M)" @@ -161,16 +169,16 @@ qed lemma (in sigma_algebra) simple_function_compose1: - assumes "simple_function f" - shows "simple_function (\x. g (f x))" + assumes "simple_function M f" + shows "simple_function M (\x. g (f x))" using simple_function_compose[OF assms, of g] by (simp add: comp_def) lemma (in sigma_algebra) simple_function_compose2: - assumes "simple_function f" and "simple_function g" - shows "simple_function (\x. h (f x) (g x))" + assumes "simple_function M f" and "simple_function M g" + shows "simple_function M (\x. h (f x) (g x))" proof - - have "simple_function ((\(x, y). h x y) \ (\x. (f x, g x)))" + have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" using assms by auto thus ?thesis by (simp_all add: comp_def) qed @@ -183,14 +191,14 @@ and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] lemma (in sigma_algebra) simple_function_setsum[intro, simp]: - assumes "\i. i \ P \ simple_function (f i)" - shows "simple_function (\x. \i\P. f i x)" + assumes "\i. i \ P \ simple_function M (f i)" + shows "simple_function M (\x. \i\P. f i x)" proof cases assume "finite P" from this assms show ?thesis by induct auto qed auto lemma (in sigma_algebra) simple_function_le_measurable: - assumes "simple_function f" "simple_function g" + assumes "simple_function M f" "simple_function M g" shows "{x \ space M. f x \ g x} \ sets M" proof - have *: "{x \ space M. f x \ g x} = @@ -214,7 +222,7 @@ lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ pextreal" assumes u: "u \ borel_measurable M" - shows "\f. (\i. simple_function (f i) \ (\x\space M. f i x \ \)) \ f \ u" + shows "\f. (\i. simple_function M (f i) \ (\x\space M. f i x \ \)) \ f \ u" proof - have "\f. \x j. (of_nat j \ u x \ f x j = j*2^j) \ (u x < of_nat j \ of_nat (f x j) \ u x * 2^j \ u x * 2^j < of_nat (Suc (f x j)))" @@ -406,10 +414,10 @@ lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': fixes u :: "'a \ pextreal" assumes "u \ borel_measurable M" - obtains (x) f where "f \ u" "\i. simple_function (f i)" "\i. \\f i`space M" + obtains (x) f where "f \ u" "\i. simple_function M (f i)" "\i. \\f i`space M" proof - from borel_measurable_implies_simple_function_sequence[OF assms] - obtain f where x: "\i. simple_function (f i)" "f \ u" + obtain f where x: "\i. simple_function M (f i)" "f \ u" and fin: "\i. \x. x\space M \ f i x \ \" by auto { fix i from fin[of _ i] have "\ \ f i`space M" by fastsimp } with x show thesis by (auto intro!: that[of f]) @@ -417,7 +425,7 @@ lemma (in sigma_algebra) simple_function_eq_borel_measurable: fixes f :: "'a \ pextreal" - shows "simple_function f \ + shows "simple_function M f \ finite (f`space M) \ f \ borel_measurable M" using simple_function_borel_measurable[of f] borel_measurable_simple_function[of f] @@ -425,8 +433,8 @@ lemma (in measure_space) simple_function_restricted: fixes f :: "'a \ pextreal" assumes "A \ sets M" - shows "sigma_algebra.simple_function (restricted_space A) f \ simple_function (\x. f x * indicator A x)" - (is "sigma_algebra.simple_function ?R f \ simple_function ?f") + shows "simple_function (restricted_space A) f \ simple_function M (\x. f x * indicator A x)" + (is "simple_function ?R f \ simple_function M ?f") proof - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) have "finite (f`A) \ finite (?f`space M)" @@ -463,29 +471,26 @@ qed lemma (in sigma_algebra) simple_function_subalgebra: - assumes "sigma_algebra.simple_function N f" - and N_subalgebra: "sets N \ sets M" "space N = space M" "sigma_algebra N" - shows "simple_function f" - using assms - unfolding simple_function_def - unfolding sigma_algebra.simple_function_def[OF N_subalgebra(3)] - by auto + assumes "simple_function N f" + and N_subalgebra: "sets N \ sets M" "space N = space M" + shows "simple_function M f" + using assms unfolding simple_function_def by auto lemma (in measure_space) simple_function_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - and f: "sigma_algebra.simple_function M' f" - shows "simple_function (\x. f (T x))" + and f: "simple_function M' f" + shows "simple_function M (\x. f (T x))" proof (intro simple_function_def[THEN iffD2] conjI ballI) interpret T: sigma_algebra M' by fact have "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto then show "finite ((\x. f (T x)) ` space M)" - using f unfolding T.simple_function_def by (auto intro: finite_subset) + using f unfolding simple_function_def by (auto intro: finite_subset) fix i assume i: "i \ (\x. f (T x)) ` space M" then have "i \ f ` space M'" using T unfolding measurable_def by auto then have "f -` {i} \ space M' \ sets M'" - using f unfolding T.simple_function_def by auto + using f unfolding simple_function_def by auto then have "T -` (f -` {i} \ space M') \ space M \ sets M" using T unfolding measurable_def by auto also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" @@ -495,12 +500,18 @@ section "Simple integral" -definition (in measure_space) simple_integral (binder "\\<^isup>S " 10) where - "simple_integral f = (\x \ f ` space M. x * \ (f -` {x} \ space M))" +definition simple_integral_def: + "integral\<^isup>S M f = (\x \ f ` space M. x * measure M (f -` {x} \ space M))" + +syntax + "_simple_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>S _. _ \_" [60,61] 110) + +translations + "\\<^isup>S x. f \M" == "CONST integral\<^isup>S M (%x. f)" lemma (in measure_space) simple_integral_cong: assumes "\t. t \ space M \ f t = g t" - shows "simple_integral f = simple_integral g" + shows "integral\<^isup>S M f = integral\<^isup>S M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" @@ -509,18 +520,18 @@ qed lemma (in measure_space) simple_integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" and "simple_function f" - shows "measure_space.simple_integral M \ f = simple_integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + and "simple_function M f" + shows "integral\<^isup>S N f = integral\<^isup>S M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact - from simple_functionD[OF `simple_function f`] assms show ?thesis - unfolding simple_integral_def v.simple_integral_def - by (auto intro!: setsum_cong) + interpret v: measure_space N + by (rule measure_space_cong) fact+ + from simple_functionD[OF `simple_function M f`] assms show ?thesis + by (auto intro!: setsum_cong simp: simple_integral_def) qed lemma (in measure_space) simple_integral_const[simp]: - "(\\<^isup>Sx. c) = c * \ (space M)" + "(\\<^isup>Sx. c \M) = c * \ (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next @@ -529,8 +540,8 @@ qed lemma (in measure_space) simple_function_partition: - assumes "simple_function f" and "simple_function g" - shows "simple_integral f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * \ A)" + assumes "simple_function M f" and "simple_function M g" + shows "integral\<^isup>S M f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * \ A)" (is "_ = setsum _ (?p ` space M)") proof- let "?sub x" = "?p ` (f -` {x} \ space M)" @@ -561,7 +572,7 @@ ultimately have "\ (f -` {f x} \ space M) = setsum (\) (?sub (f x))" by (subst measure_finitely_additive) auto } - hence "simple_integral f = (\(x,A)\?SIGMA. x * \ A)" + hence "integral\<^isup>S M f = (\(x,A)\?SIGMA. x * \ A)" unfolding simple_integral_def by (subst setsum_Sigma[symmetric], auto intro!: setsum_cong simp: setsum_right_distrib[symmetric]) @@ -584,8 +595,8 @@ qed lemma (in measure_space) simple_integral_add[simp]: - assumes "simple_function f" and "simple_function g" - shows "(\\<^isup>Sx. f x + g x) = simple_integral f + simple_integral g" + assumes "simple_function M f" and "simple_function M g" + shows "(\\<^isup>Sx. f x + g x \M) = integral\<^isup>S M f + integral\<^isup>S M g" proof - { fix x let ?S = "g -` {g x} \ f -` {f x} \ space M" assume "x \ space M" @@ -595,15 +606,15 @@ thus ?thesis unfolding simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]] - simple_function_partition[OF `simple_function f` `simple_function g`] - simple_function_partition[OF `simple_function g` `simple_function f`] + simple_function_partition[OF `simple_function M f` `simple_function M g`] + simple_function_partition[OF `simple_function M g` `simple_function M f`] apply (subst (3) Int_commute) by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong) qed lemma (in measure_space) simple_integral_setsum[simp]: - assumes "\i. i \ P \ simple_function (f i)" - shows "(\\<^isup>Sx. \i\P. f i x) = (\i\P. simple_integral (f i))" + assumes "\i. i \ P \ simple_function M (f i)" + shows "(\\<^isup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^isup>S M (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -611,8 +622,8 @@ qed auto lemma (in measure_space) simple_integral_mult[simp]: - assumes "simple_function f" - shows "(\\<^isup>Sx. c * f x) = c * simple_integral f" + assumes "simple_function M f" + shows "(\\<^isup>Sx. c * f x \M) = c * integral\<^isup>S M f" proof - note mult = simple_function_mult[OF simple_function_const[of c] assms] { fix x let ?S = "f -` {f x} \ (\x. c * f x) -` {c * f x} \ space M" @@ -626,8 +637,8 @@ qed lemma (in sigma_algebra) simple_function_If: - assumes sf: "simple_function f" "simple_function g" and A: "A \ sets M" - shows "simple_function (\x. if x \ A then f x else g x)" (is "simple_function ?IF") + assumes sf: "simple_function M f" "simple_function M g" and A: "A \ sets M" + shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") proof - def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" show ?thesis unfolding simple_function_def @@ -648,17 +659,17 @@ qed lemma (in measure_space) simple_integral_mono_AE: - assumes "simple_function f" and "simple_function g" + assumes "simple_function M f" and "simple_function M g" and mono: "AE x. f x \ g x" - shows "simple_integral f \ simple_integral g" + shows "integral\<^isup>S M f \ integral\<^isup>S M g" proof - let "?S x" = "(g -` {g x} \ space M) \ (f -` {f x} \ space M)" have *: "\x. g -` {g x} \ f -` {f x} \ space M = ?S x" "\x. f -` {f x} \ g -` {g x} \ space M = ?S x" by auto show ?thesis unfolding * - simple_function_partition[OF `simple_function f` `simple_function g`] - simple_function_partition[OF `simple_function g` `simple_function f`] + simple_function_partition[OF `simple_function M f` `simple_function M g`] + simple_function_partition[OF `simple_function M g` `simple_function M f`] proof (safe intro!: setsum_mono) fix x assume "x \ space M" then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto @@ -680,23 +691,23 @@ qed lemma (in measure_space) simple_integral_mono: - assumes "simple_function f" and "simple_function g" + assumes "simple_function M f" and "simple_function M g" and mono: "\ x. x \ space M \ f x \ g x" - shows "simple_integral f \ simple_integral g" + shows "integral\<^isup>S M f \ integral\<^isup>S M g" proof (rule simple_integral_mono_AE[OF assms(1, 2)]) show "AE x. f x \ g x" using mono by (rule AE_cong) auto qed lemma (in measure_space) simple_integral_cong_AE: - assumes "simple_function f" "simple_function g" and "AE x. f x = g x" - shows "simple_integral f = simple_integral g" + assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x" + shows "integral\<^isup>S M f = integral\<^isup>S M g" using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) lemma (in measure_space) simple_integral_cong': - assumes sf: "simple_function f" "simple_function g" + assumes sf: "simple_function M f" "simple_function M g" and mea: "\ {x\space M. f x \ g x} = 0" - shows "simple_integral f = simple_integral g" + shows "integral\<^isup>S M f = integral\<^isup>S M g" proof (intro simple_integral_cong_AE sf AE_I) show "\ {x\space M. f x \ g x} = 0" by fact show "{x \ space M. f x \ g x} \ sets M" @@ -705,12 +716,12 @@ lemma (in measure_space) simple_integral_indicator: assumes "A \ sets M" - assumes "simple_function f" - shows "(\\<^isup>Sx. f x * indicator A x) = + assumes "simple_function M f" + shows "(\\<^isup>Sx. f x * indicator A x \M) = (\x \ f ` space M. x * \ (f -` {x} \ space M \ A))" proof cases assume "A = space M" - moreover hence "(\\<^isup>Sx. f x * indicator A x) = simple_integral f" + moreover hence "(\\<^isup>Sx. f x * indicator A x \M) = integral\<^isup>S M f" by (auto intro!: simple_integral_cong) moreover have "\X. X \ space M \ space M = X \ space M" by auto ultimately show ?thesis by (simp add: simple_integral_def) @@ -726,7 +737,7 @@ next show "0 \ ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x]) qed - have *: "(\\<^isup>Sx. f x * indicator A x) = + have *: "(\\<^isup>Sx. f x * indicator A x \M) = (\x \ f ` space M \ {0}. x * \ (f -` {x} \ space M \ A))" unfolding simple_integral_def I proof (rule setsum_mono_zero_cong_left) @@ -752,7 +763,7 @@ lemma (in measure_space) simple_integral_indicator_only[simp]: assumes "A \ sets M" - shows "simple_integral (indicator A) = \ A" + shows "integral\<^isup>S M (indicator A) = \ A" proof cases assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto thus ?thesis unfolding simple_integral_def using `space M = {}` by auto @@ -765,22 +776,22 @@ qed lemma (in measure_space) simple_integral_null_set: - assumes "simple_function u" "N \ null_sets" - shows "(\\<^isup>Sx. u x * indicator N x) = 0" + assumes "simple_function M u" "N \ null_sets" + shows "(\\<^isup>Sx. u x * indicator N x \M) = 0" proof - have "AE x. indicator N x = (0 :: pextreal)" using `N \ null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N]) - then have "(\\<^isup>Sx. u x * indicator N x) = (\\<^isup>Sx. 0)" + then have "(\\<^isup>Sx. u x * indicator N x \M) = (\\<^isup>Sx. 0 \M)" using assms by (intro simple_integral_cong_AE) (auto intro!: AE_disjI2) then show ?thesis by simp qed lemma (in measure_space) simple_integral_cong_AE_mult_indicator: - assumes sf: "simple_function f" and eq: "AE x. x \ S" and "S \ sets M" - shows "simple_integral f = (\\<^isup>Sx. f x * indicator S x)" + assumes sf: "simple_function M f" and eq: "AE x. x \ S" and "S \ sets M" + shows "integral\<^isup>S M f = (\\<^isup>Sx. f x * indicator S x \M)" proof (rule simple_integral_cong_AE) - show "simple_function f" by fact - show "simple_function (\x. f x * indicator S x)" + show "simple_function M f" by fact + show "simple_function M (\x. f x * indicator S x)" using sf `S \ sets M` by auto from eq show "AE x. f x = f x * indicator S x" by (rule AE_mp) simp @@ -788,10 +799,9 @@ lemma (in measure_space) simple_integral_restricted: assumes "A \ sets M" - assumes sf: "simple_function (\x. f x * indicator A x)" - shows "measure_space.simple_integral (restricted_space A) \ f = (\\<^isup>Sx. f x * indicator A x)" - (is "_ = simple_integral ?f") - unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \ sets M`]] + assumes sf: "simple_function M (\x. f x * indicator A x)" + shows "integral\<^isup>S (restricted_space A) f = (\\<^isup>Sx. f x * indicator A x \M)" + (is "_ = integral\<^isup>S M ?f") unfolding simple_integral_def proof (simp, safe intro!: setsum_mono_zero_cong_left) from sf show "finite (?f ` space M)" @@ -816,64 +826,71 @@ qed lemma (in measure_space) simple_integral_subalgebra: - assumes N: "measure_space N \" and [simp]: "space N = space M" - shows "measure_space.simple_integral N \ = simple_integral" - unfolding simple_integral_def_raw - unfolding measure_space.simple_integral_def_raw[OF N] by simp + assumes N: "measure_space N" and [simp]: "space N = space M" "measure N = measure M" + shows "integral\<^isup>S N = integral\<^isup>S M" + unfolding simple_integral_def_raw by simp lemma (in measure_space) simple_integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - and f: "sigma_algebra.simple_function M' f" - shows "measure_space.simple_integral M' (\A. \ (T -` A \ space M)) f = (\\<^isup>S x. f (T x))" - (is "measure_space.simple_integral M' ?nu f = _") + and f: "simple_function M' f" + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + shows "integral\<^isup>S M' f = (\\<^isup>S x. f (T x) \M)" proof - - interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto - show "T.simple_integral f = (\\<^isup>S x. f (T x))" - unfolding simple_integral_def T.simple_integral_def + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) + show "integral\<^isup>S M' f = (\\<^isup>S x. f (T x) \M)" + unfolding simple_integral_def proof (intro setsum_mono_zero_cong_right ballI) show "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto show "finite (f ` space M')" - using f unfolding T.simple_function_def by auto + using f unfolding simple_function_def by auto next fix i assume "i \ f ` space M' - (\x. f (T x)) ` space M" then have "T -` (f -` {i} \ space M') \ space M = {}" by (auto simp: image_iff) - then show "i * \ (T -` (f -` {i} \ space M') \ space M) = 0" by simp + with f[THEN T.simple_functionD(2), THEN \] + show "i * T.\ (f -` {i} \ space M') = 0" by simp next fix i assume "i \ (\x. f (T x)) ` space M" then have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" using T unfolding measurable_def by auto - then show "i * \ (T -` (f -` {i} \ space M') \ space M) = i * \ ((\x. f (T x)) -` {i} \ space M)" + with f[THEN T.simple_functionD(2), THEN \] + show "i * T.\ (f -` {i} \ space M') = i * \ ((\x. f (T x)) -` {i} \ space M)" by auto qed qed -section "Continuous posititve integration" +section "Continuous positive integration" + +definition positive_integral_def: + "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ f}. integral\<^isup>S M g)" -definition (in measure_space) positive_integral (binder "\\<^isup>+ " 10) where - "positive_integral f = SUPR {g. simple_function g \ g \ f} simple_integral" +syntax + "_positive_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>+ _. _ \_" [60,61] 110) + +translations + "\\<^isup>+ x. f \M" == "CONST integral\<^isup>P M (%x. f)" -lemma (in measure_space) positive_integral_alt: - "positive_integral f = - (SUPR {g. simple_function g \ g \ f \ \ \ g`space M} simple_integral)" (is "_ = ?alt") +lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f = + (SUP g : {g. simple_function M g \ g \ f \ \ \ g`space M}. integral\<^isup>S M g)" + (is "_ = ?alt") proof (rule antisym SUP_leI) - show "positive_integral f \ ?alt" unfolding positive_integral_def + show "integral\<^isup>P M f \ ?alt" unfolding positive_integral_def proof (safe intro!: SUP_leI) - fix g assume g: "simple_function g" "g \ f" + fix g assume g: "simple_function M g" "g \ f" let ?G = "g -` {\} \ space M" - show "simple_integral g \ - SUPR {g. simple_function g \ g \ f \ \ \ g ` space M} simple_integral" - (is "simple_integral g \ SUPR ?A simple_integral") + show "integral\<^isup>S M g \ + (SUP h : {i. simple_function M i \ i \ f \ \ \ i ` space M}. integral\<^isup>S M h)" + (is "integral\<^isup>S M g \ SUPR ?A _") proof cases let ?g = "\x. indicator (space M - ?G) x * g x" - have g': "simple_function ?g" + have g': "simple_function M ?g" using g by (auto intro: simple_functionD) moreover assume "\ ?G = 0" then have "AE x. g x = ?g x" using g by (intro AE_I[where N="?G"]) (auto intro: simple_functionD simp: indicator_def) - with g(1) g' have "simple_integral g = simple_integral ?g" + with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g" by (rule simple_integral_cong_AE) moreover have "?g \ g" by (auto simp: le_fun_def indicator_def) from this `g \ f` have "?g \ f" by (rule order_trans) @@ -885,15 +902,15 @@ then have "?G \ {}" by auto then have "\ \ g`space M" by force then have "space M \ {}" by auto - have "SUPR ?A simple_integral = \" + have "SUPR ?A (integral\<^isup>S M) = \" proof (intro SUP_\[THEN iffD2] allI impI) fix x assume "x < \" then guess n unfolding less_\_Ex_of_nat .. note n = this then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp let ?g = "\x. (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * indicator ?G x" - show "\i\?A. x < simple_integral i" + show "\i\?A. x < integral\<^isup>S M i" proof (intro bexI impI CollectI conjI) - show "simple_function ?g" using g + show "simple_function M ?g" using g by (auto intro!: simple_functionD simple_function_add) have "?g \ g" by (auto simp: le_fun_def indicator_def) from this g(2) show "?g \ f" by (rule order_trans) @@ -902,10 +919,10 @@ have "x < (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * \ ?G" using n `\ ?G \ 0` `0 < n` by (auto simp: pextreal_noteq_omega_Ex field_simps) - also have "\ = simple_integral ?g" using g `space M \ {}` + also have "\ = integral\<^isup>S M ?g" using g `space M \ {}` by (subst simple_integral_indicator) (auto simp: image_constant ac_simps dest: simple_functionD) - finally show "x < simple_integral ?g" . + finally show "x < integral\<^isup>S M ?g" . qed qed then show ?thesis by simp @@ -914,40 +931,41 @@ qed (auto intro!: SUP_subset simp: positive_integral_def) lemma (in measure_space) positive_integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" - shows "measure_space.positive_integral M \ f = positive_integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact + interpret v: measure_space N + by (rule measure_space_cong) fact+ with assms show ?thesis - unfolding positive_integral_def v.positive_integral_def SUPR_def + unfolding positive_integral_def SUPR_def by (auto intro!: arg_cong[where f=Sup] image_cong - simp: simple_integral_cong_measure[of \]) + simp: simple_integral_cong_measure[OF assms] + simple_function_cong_algebra[OF assms(2,3)]) qed lemma (in measure_space) positive_integral_alt1: - "positive_integral f = - (SUP g : {g. simple_function g \ (\x\space M. g x \ f x \ g x \ \)}. simple_integral g)" + "integral\<^isup>P M f = + (SUP g : {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}. integral\<^isup>S M g)" unfolding positive_integral_alt SUPR_def proof (safe intro!: arg_cong[where f=Sup]) fix g let ?g = "\x. if x \ space M then g x else f x" - assume "simple_function g" "\x\space M. g x \ f x \ g x \ \" - hence "?g \ f" "simple_function ?g" "simple_integral g = simple_integral ?g" + assume "simple_function M g" "\x\space M. g x \ f x \ g x \ \" + hence "?g \ f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g" "\ \ g`space M" unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong) - thus "simple_integral g \ simple_integral ` {g. simple_function g \ g \ f \ \ \ g`space M}" + thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ g \ f \ \ \ g`space M}" by auto next - fix g assume "simple_function g" "g \ f" "\ \ g`space M" - hence "simple_function g" "\x\space M. g x \ f x \ g x \ \" + fix g assume "simple_function M g" "g \ f" "\ \ g`space M" + hence "simple_function M g" "\x\space M. g x \ f x \ g x \ \" by (auto simp add: le_fun_def image_iff) - thus "simple_integral g \ simple_integral ` {g. simple_function g \ (\x\space M. g x \ f x \ g x \ \)}" + thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}" by auto qed lemma (in measure_space) positive_integral_cong: assumes "\x. x \ space M \ f x = g x" - shows "positive_integral f = positive_integral g" + shows "integral\<^isup>P M f = integral\<^isup>P M g" proof - have "\h. (\x\space M. h x \ f x \ h x \ \) = (\x\space M. h x \ g x \ h x \ \)" using assms by auto @@ -955,30 +973,30 @@ qed lemma (in measure_space) positive_integral_eq_simple_integral: - assumes "simple_function f" - shows "positive_integral f = simple_integral f" + assumes "simple_function M f" + shows "integral\<^isup>P M f = integral\<^isup>S M f" unfolding positive_integral_def proof (safe intro!: pextreal_SUPI) - fix g assume "simple_function g" "g \ f" - with assms show "simple_integral g \ simple_integral f" + fix g assume "simple_function M g" "g \ f" + with assms show "integral\<^isup>S M g \ integral\<^isup>S M f" by (auto intro!: simple_integral_mono simp: le_fun_def) next - fix y assume "\x. x\{g. simple_function g \ g \ f} \ simple_integral x \ y" - with assms show "simple_integral f \ y" by auto + fix y assume "\x. x\{g. simple_function M g \ g \ f} \ integral\<^isup>S M x \ y" + with assms show "integral\<^isup>S M f \ y" by auto qed lemma (in measure_space) positive_integral_mono_AE: assumes ae: "AE x. u x \ v x" - shows "positive_integral u \ positive_integral v" + shows "integral\<^isup>P M u \ integral\<^isup>P M v" unfolding positive_integral_alt1 proof (safe intro!: SUPR_mono) - fix a assume a: "simple_function a" and mono: "\x\space M. a x \ u x \ a x \ \" + fix a assume a: "simple_function M a" and mono: "\x\space M. a x \ u x \ a x \ \" from ae obtain N where N: "{x\space M. \ u x \ v x} \ N" "N \ sets M" "\ N = 0" by (auto elim!: AE_E) - have "simple_function (\x. a x * indicator (space M - N) x)" + have "simple_function M (\x. a x * indicator (space M - N) x)" using `N \ sets M` a by auto - with a show "\b\{g. simple_function g \ (\x\space M. g x \ v x \ g x \ \)}. - simple_integral a \ simple_integral b" + with a show "\b\{g. simple_function M g \ (\x\space M. g x \ v x \ g x \ \)}. + integral\<^isup>S M a \ integral\<^isup>S M b" proof (safe intro!: bexI[of _ "\x. a x * indicator (space M - N) x"] simple_integral_mono_AE) show "AE x. a x \ a x * indicator (space M - N) x" @@ -987,7 +1005,7 @@ N \ {x \ space M. a x \ 0}" (is "?N = _") using `N \ sets M`[THEN sets_into_space] by (auto simp: indicator_def) then show "?N \ sets M" - using `N \ sets M` `simple_function a`[THEN borel_measurable_simple_function] + using `N \ sets M` `simple_function M a`[THEN borel_measurable_simple_function] by (auto intro!: measure_mono Int) then have "\ ?N \ \ N" unfolding * using `N \ sets M` by (auto intro!: measure_mono) @@ -1010,12 +1028,12 @@ qed lemma (in measure_space) positive_integral_cong_AE: - "AE x. u x = v x \ positive_integral u = positive_integral v" + "AE x. u x = v x \ integral\<^isup>P M u = integral\<^isup>P M v" by (auto simp: eq_iff intro!: positive_integral_mono_AE) lemma (in measure_space) positive_integral_mono: assumes mono: "\x. x \ space M \ u x \ v x" - shows "positive_integral u \ positive_integral v" + shows "integral\<^isup>P M u \ integral\<^isup>P M v" using mono by (auto intro!: AE_cong positive_integral_mono_AE) lemma image_set_cong: @@ -1027,15 +1045,15 @@ lemma (in measure_space) positive_integral_SUP_approx: assumes "f \ s" and f: "\i. f i \ borel_measurable M" - and "simple_function u" + and "simple_function M u" and le: "u \ s" and real: "\ \ u`space M" - shows "simple_integral u \ (SUP i. positive_integral (f i))" (is "_ \ ?S") + shows "integral\<^isup>S M u \ (SUP i. integral\<^isup>P M (f i))" (is "_ \ ?S") proof (rule pextreal_le_mult_one_interval) fix a :: pextreal assume "0 < a" "a < 1" hence "a \ 0" by auto let "?B i" = "{x \ space M. a * u x \ f i x}" have B: "\i. ?B i \ sets M" - using f `simple_function u` by (auto simp: borel_measurable_simple_function) + using f `simple_function M u` by (auto simp: borel_measurable_simple_function) let "?uB i x" = "u x * indicator (?B i) x" @@ -1049,7 +1067,7 @@ note B_mono = this have u: "\x. x \ space M \ u -` {u x} \ space M \ sets M" - using `simple_function u` by (auto simp add: simple_function_def) + using `simple_function M u` by (auto simp add: simple_function_def) have "\i. (\n. (u -` {i} \ space M) \ ?B n) \ (u -` {i} \ space M)" using B_mono unfolding isoton_def proof safe @@ -1071,8 +1089,8 @@ qed auto note measure_conv = measure_up[OF Int[OF u B] this] - have "simple_integral u = (SUP i. simple_integral (?uB i))" - unfolding simple_integral_indicator[OF B `simple_function u`] + have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))" + unfolding simple_integral_indicator[OF B `simple_function M u`] proof (subst SUPR_pextreal_setsum, safe) fix x n assume "x \ space M" have "\ (u -` {u x} \ space M \ {x \ space M. a * u x \ f n x}) @@ -1082,52 +1100,51 @@ \ u x * \ (u -` {u x} \ space M \ ?B (Suc n))" by (auto intro: mult_left_mono) next - show "simple_integral u = + show "integral\<^isup>S M u = (\i\u ` space M. SUP n. i * \ (u -` {i} \ space M \ ?B n))" using measure_conv unfolding simple_integral_def isoton_def by (auto intro!: setsum_cong simp: pextreal_SUP_cmult) qed moreover - have "a * (SUP i. simple_integral (?uB i)) \ ?S" + have "a * (SUP i. integral\<^isup>S M (?uB i)) \ ?S" unfolding pextreal_SUP_cmult[symmetric] proof (safe intro!: SUP_mono bexI) fix i - have "a * simple_integral (?uB i) = (\\<^isup>Sx. a * ?uB i x)" - using B `simple_function u` + have "a * integral\<^isup>S M (?uB i) = (\\<^isup>Sx. a * ?uB i x \M)" + using B `simple_function M u` by (subst simple_integral_mult[symmetric]) (auto simp: field_simps) - also have "\ \ positive_integral (f i)" + also have "\ \ integral\<^isup>P M (f i)" proof - have "\x. a * ?uB i x \ f i x" unfolding indicator_def by auto - hence *: "simple_function (\x. a * ?uB i x)" using B assms(3) + hence *: "simple_function M (\x. a * ?uB i x)" using B assms(3) by (auto intro!: simple_integral_mono) show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric] by (auto intro!: positive_integral_mono simp: indicator_def) qed - finally show "a * simple_integral (?uB i) \ positive_integral (f i)" + finally show "a * integral\<^isup>S M (?uB i) \ integral\<^isup>P M (f i)" by auto qed simp - ultimately show "a * simple_integral u \ ?S" by simp + ultimately show "a * integral\<^isup>S M u \ ?S" by simp qed text {* Beppo-Levi monotone convergence theorem *} lemma (in measure_space) positive_integral_isoton: assumes "f \ u" "\i. f i \ borel_measurable M" - shows "(\i. positive_integral (f i)) \ positive_integral u" + shows "(\i. integral\<^isup>P M (f i)) \ integral\<^isup>P M u" unfolding isoton_def proof safe - fix i show "positive_integral (f i) \ positive_integral (f (Suc i))" + fix i show "integral\<^isup>P M (f i) \ integral\<^isup>P M (f (Suc i))" apply (rule positive_integral_mono) using `f \ u` unfolding isoton_def le_fun_def by auto next have u: "u = (SUP i. f i)" using `f \ u` unfolding isoton_def by auto - - show "(SUP i. positive_integral (f i)) = positive_integral u" + show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u" proof (rule antisym) from `f \ u`[THEN isoton_Sup, unfolded le_fun_def] - show "(SUP j. positive_integral (f j)) \ positive_integral u" + show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M u" by (auto intro!: SUP_leI positive_integral_mono) next - show "positive_integral u \ (SUP i. positive_integral (f i))" + show "integral\<^isup>P M u \ (SUP i. integral\<^isup>P M (f i))" unfolding positive_integral_alt[of u] by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms]) qed @@ -1136,12 +1153,12 @@ lemma (in measure_space) positive_integral_monotone_convergence_SUP: assumes "\i x. x \ space M \ f i x \ f (Suc i) x" assumes "\i. f i \ borel_measurable M" - shows "(SUP i. positive_integral (f i)) = (\\<^isup>+ x. SUP i. f i x)" - (is "_ = positive_integral ?u") + shows "(SUP i. integral\<^isup>P M (f i)) = (\\<^isup>+ x. (SUP i. f i x) \M)" + (is "_ = integral\<^isup>P M ?u") proof - show ?thesis proof (rule antisym) - show "(SUP j. positive_integral (f j)) \ positive_integral ?u" + show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M ?u" by (auto intro!: SUP_leI positive_integral_mono le_SUPI) next def rf \ "\i. \x\space M. f i x" and ru \ "\x\space M. ?u x" @@ -1151,26 +1168,26 @@ unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply using SUP_const[OF UNIV_not_empty] by (auto simp: restrict_def le_fun_def fun_eq_iff) - ultimately have "positive_integral ru \ (SUP i. positive_integral (rf i))" + ultimately have "integral\<^isup>P M ru \ (SUP i. integral\<^isup>P M (rf i))" unfolding positive_integral_alt[of ru] by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx) - then show "positive_integral ?u \ (SUP i. positive_integral (f i))" + then show "integral\<^isup>P M ?u \ (SUP i. integral\<^isup>P M (f i))" unfolding ru_def rf_def by (simp cong: positive_integral_cong) qed qed lemma (in measure_space) SUP_simple_integral_sequences: - assumes f: "f \ u" "\i. simple_function (f i)" - and g: "g \ u" "\i. simple_function (g i)" - shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))" + assumes f: "f \ u" "\i. simple_function M (f i)" + and g: "g \ u" "\i. simple_function M (g i)" + shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))" (is "SUPR _ ?F = SUPR _ ?G") proof - - have "(SUP i. ?F i) = (SUP i. positive_integral (f i))" + have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))" using assms by (simp add: positive_integral_eq_simple_integral) - also have "\ = positive_integral u" + also have "\ = integral\<^isup>P M u" using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]] unfolding isoton_def by simp - also have "\ = (SUP i. positive_integral (g i))" + also have "\ = (SUP i. integral\<^isup>P M (g i))" using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]] unfolding isoton_def by simp also have "\ = (SUP i. ?G i)" @@ -1179,38 +1196,36 @@ qed lemma (in measure_space) positive_integral_const[simp]: - "(\\<^isup>+ x. c) = c * \ (space M)" + "(\\<^isup>+ x. c \M) = c * \ (space M)" by (subst positive_integral_eq_simple_integral) auto lemma (in measure_space) positive_integral_isoton_simple: - assumes "f \ u" and e: "\i. simple_function (f i)" - shows "(\i. simple_integral (f i)) \ positive_integral u" + assumes "f \ u" and e: "\i. simple_function M (f i)" + shows "(\i. integral\<^isup>S M (f i)) \ integral\<^isup>P M u" using positive_integral_isoton[OF `f \ u` e(1)[THEN borel_measurable_simple_function]] unfolding positive_integral_eq_simple_integral[OF e] . lemma (in measure_space) positive_integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" and f: "f \ borel_measurable M'" - shows "measure_space.positive_integral M' (\A. \ (T -` A \ space M)) f = (\\<^isup>+ x. f (T x))" - (is "measure_space.positive_integral M' ?nu f = _") + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + shows "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" proof - - interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto - obtain f' where f': "f' \ f" "\i. T.simple_function (f' i)" + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) + obtain f' where f': "f' \ f" "\i. simple_function M' (f' i)" using T.borel_measurable_implies_simple_function_sequence[OF f] by blast - then have f: "(\i x. f' i (T x)) \ (\x. f (T x))" "\i. simple_function (\x. f' i (T x))" + then have f: "(\i x. f' i (T x)) \ (\x. f (T x))" "\i. simple_function M (\x. f' i (T x))" using simple_function_vimage[OF T] unfolding isoton_fun_expand by auto - show "T.positive_integral f = (\\<^isup>+ x. f (T x))" + show "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" using positive_integral_isoton_simple[OF f] using T.positive_integral_isoton_simple[OF f'] - unfolding simple_integral_vimage[OF T f'(2)] isoton_def - by simp + by (simp add: simple_integral_vimage[OF T f'(2) \] isoton_def) qed lemma (in measure_space) positive_integral_linear: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - shows "(\\<^isup>+ x. a * f x + g x) = - a * positive_integral f + positive_integral g" - (is "positive_integral ?L = _") + shows "(\\<^isup>+ x. a * f x + g x \M) = a * integral\<^isup>P M f + integral\<^isup>P M g" + (is "integral\<^isup>P M ?L = _") proof - from borel_measurable_implies_simple_function_sequence'[OF f] guess u . note u = this positive_integral_isoton_simple[OF this(1-2)] @@ -1222,46 +1237,45 @@ using assms by simp from borel_measurable_implies_simple_function_sequence'[OF this] guess l . note positive_integral_isoton_simple[OF this(1-2)] and l = this - moreover have - "(SUP i. simple_integral (l i)) = (SUP i. simple_integral (?L' i))" + moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))" proof (rule SUP_simple_integral_sequences[OF l(1-2)]) - show "?L' \ ?L" "\i. simple_function (?L' i)" + show "?L' \ ?L" "\i. simple_function M (?L' i)" using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right) qed moreover from u v have L'_isoton: - "(\i. simple_integral (?L' i)) \ a * positive_integral f + positive_integral g" + "(\i. integral\<^isup>S M (?L' i)) \ a * integral\<^isup>P M f + integral\<^isup>P M g" by (simp add: isoton_add isoton_cmult_right) ultimately show ?thesis by (simp add: isoton_def) qed lemma (in measure_space) positive_integral_cmult: assumes "f \ borel_measurable M" - shows "(\\<^isup>+ x. c * f x) = c * positive_integral f" + shows "(\\<^isup>+ x. c * f x \M) = c * integral\<^isup>P M f" using positive_integral_linear[OF assms, of "\x. 0" c] by auto lemma (in measure_space) positive_integral_multc: assumes "f \ borel_measurable M" - shows "(\\<^isup>+ x. f x * c) = positive_integral f * c" + shows "(\\<^isup>+ x. f x * c \M) = integral\<^isup>P M f * c" unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp lemma (in measure_space) positive_integral_indicator[simp]: - "A \ sets M \ (\\<^isup>+ x. indicator A x) = \ A" + "A \ sets M \ (\\<^isup>+ x. indicator A x\M) = \ A" by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_cmult_indicator: - "A \ sets M \ (\\<^isup>+ x. c * indicator A x) = c * \ A" + "A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * \ A" by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_add: assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "(\\<^isup>+ x. f x + g x) = positive_integral f + positive_integral g" + shows "(\\<^isup>+ x. f x + g x \M) = integral\<^isup>P M f + integral\<^isup>P M g" using positive_integral_linear[OF assms, of 1] by simp lemma (in measure_space) positive_integral_setsum: assumes "\i. i\P \ f i \ borel_measurable M" - shows "(\\<^isup>+ x. \i\P. f i x) = (\i\P. positive_integral (f i))" + shows "(\\<^isup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^isup>P M (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -1277,14 +1291,13 @@ lemma (in measure_space) positive_integral_diff: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - and fin: "positive_integral g \ \" + and fin: "integral\<^isup>P M g \ \" and mono: "\x. x \ space M \ g x \ f x" - shows "(\\<^isup>+ x. f x - g x) = positive_integral f - positive_integral g" + shows "(\\<^isup>+ x. f x - g x \M) = integral\<^isup>P M f - integral\<^isup>P M g" proof - have borel: "(\x. f x - g x) \ borel_measurable M" using f g by (rule borel_measurable_pextreal_diff) - have "(\\<^isup>+x. f x - g x) + positive_integral g = - positive_integral f" + have "(\\<^isup>+x. f x - g x \M) + integral\<^isup>P M g = integral\<^isup>P M f" unfolding positive_integral_add[OF borel g, symmetric] proof (rule positive_integral_cong) fix x assume "x \ space M" @@ -1297,9 +1310,9 @@ lemma (in measure_space) positive_integral_psuminf: assumes "\i. f i \ borel_measurable M" - shows "(\\<^isup>+ x. \\<^isub>\ i. f i x) = (\\<^isub>\ i. positive_integral (f i))" + shows "(\\<^isup>+ x. (\\<^isub>\ i. f i x) \M) = (\\<^isub>\ i. integral\<^isup>P M (f i))" proof - - have "(\i. (\\<^isup>+x. \i (\\<^isup>+x. \\<^isub>\i. f i x)" + have "(\i. (\\<^isup>+x. (\iM)) \ (\\<^isup>+x. (\\<^isub>\i. f i x) \M)" by (rule positive_integral_isoton) (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono arg_cong[where f=Sup] @@ -1312,79 +1325,86 @@ lemma (in measure_space) positive_integral_lim_INF: fixes u :: "nat \ 'a \ pextreal" assumes "\i. u i \ borel_measurable M" - shows "(\\<^isup>+ x. SUP n. INF m. u (m + n) x) \ - (SUP n. INF m. positive_integral (u (m + n)))" + shows "(\\<^isup>+ x. (SUP n. INF m. u (m + n) x) \M) \ + (SUP n. INF m. integral\<^isup>P M (u (m + n)))" proof - - have "(\\<^isup>+x. SUP n. INF m. u (m + n) x) - = (SUP n. (\\<^isup>+x. INF m. u (m + n) x))" + have "(\\<^isup>+x. (SUP n. INF m. u (m + n) x) \M) + = (SUP n. (\\<^isup>+x. (INF m. u (m + n) x) \M))" using assms by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono) (auto simp del: add_Suc simp add: add_Suc[symmetric]) - also have "\ \ (SUP n. INF m. positive_integral (u (m + n)))" + also have "\ \ (SUP n. INF m. integral\<^isup>P M (u (m + n)))" by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI) finally show ?thesis . qed lemma (in measure_space) measure_space_density: assumes borel: "u \ borel_measurable M" - shows "measure_space M (\A. (\\<^isup>+ x. u x * indicator A x))" (is "measure_space M ?v") -proof - show "?v {} = 0" by simp - show "countably_additive M ?v" - unfolding countably_additive_def - proof safe - fix A :: "nat \ 'a set" - assume "range A \ sets M" - hence "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" - using borel by (auto intro: borel_measurable_indicator) - moreover assume "disjoint_family A" - note psuminf_indicator[OF this] - ultimately show "(\\<^isub>\n. ?v (A n)) = ?v (\x. A x)" - by (simp add: positive_integral_psuminf[symmetric]) + and M'[simp]: "M' = (M\measure := \A. (\\<^isup>+ x. u x * indicator A x \M)\)" + shows "measure_space M'" +proof - + interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto + show ?thesis + proof + show "measure M' {} = 0" unfolding M' by simp + show "countably_additive M' (measure M')" + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume "range A \ sets M'" + then have "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" + using borel by (auto intro: borel_measurable_indicator) + moreover assume "disjoint_family A" + note psuminf_indicator[OF this] + ultimately show "(\\<^isub>\n. measure M' (A n)) = measure M' (\x. A x)" + by (simp add: positive_integral_psuminf[symmetric]) + qed qed qed lemma (in measure_space) positive_integral_translated_density: assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "measure_space.positive_integral M (\A. (\\<^isup>+ x. f x * indicator A x)) g = - (\\<^isup>+ x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") + and M': "M' = (M\ measure := \A. (\\<^isup>+ x. f x * indicator A x \M)\)" + shows "integral\<^isup>P M' g = (\\<^isup>+ x. f x * g x \M)" proof - - from measure_space_density[OF assms(1)] - interpret T: measure_space M ?T . + from measure_space_density[OF assms(1) M'] + interpret T: measure_space M' . + have borel[simp]: + "borel_measurable M' = borel_measurable M" + "simple_function M' = simple_function M" + unfolding measurable_def simple_function_def_raw by (auto simp: M') from borel_measurable_implies_simple_function_sequence[OF assms(2)] - obtain G where G: "\i. simple_function (G i)" "G \ g" by blast + obtain G where G: "\i. simple_function M (G i)" "G \ g" by blast note G_borel = borel_measurable_simple_function[OF this(1)] - from T.positive_integral_isoton[OF `G \ g` G_borel] - have *: "(\i. T.positive_integral (G i)) \ T.positive_integral g" . + from T.positive_integral_isoton[unfolded borel, OF `G \ g` G_borel] + have *: "(\i. integral\<^isup>P M' (G i)) \ integral\<^isup>P M' g" . { fix i have [simp]: "finite (G i ` space M)" using G(1) unfolding simple_function_def by auto - have "T.positive_integral (G i) = T.simple_integral (G i)" + have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)" using G T.positive_integral_eq_simple_integral by simp - also have "\ = (\\<^isup>+x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x))" - apply (simp add: T.simple_integral_def) + also have "\ = (\\<^isup>+x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x) \M)" + apply (simp add: simple_integral_def M') apply (subst positive_integral_cmult[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) apply (subst positive_integral_setsum[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) by (simp add: setsum_right_distrib field_simps) - also have "\ = (\\<^isup>+x. f x * G i x)" + also have "\ = (\\<^isup>+x. f x * G i x \M)" by (auto intro!: positive_integral_cong simp: indicator_def if_distrib setsum_cases) - finally have "T.positive_integral (G i) = (\\<^isup>+x. f x * G i x)" . } - with * have eq_Tg: "(\i. (\\<^isup>+x. f x * G i x)) \ T.positive_integral g" by simp + finally have "integral\<^isup>P M' (G i) = (\\<^isup>+x. f x * G i x \M)" . } + with * have eq_Tg: "(\i. (\\<^isup>+x. f x * G i x \M)) \ integral\<^isup>P M' g" by simp from G(2) have "(\i x. f x * G i x) \ (\x. f x * g x)" unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) - then have "(\i. (\\<^isup>+x. f x * G i x)) \ (\\<^isup>+x. f x * g x)" + then have "(\i. (\\<^isup>+x. f x * G i x \M)) \ (\\<^isup>+x. f x * g x \M)" using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times) - with eq_Tg show "T.positive_integral g = (\\<^isup>+x. f x * g x)" + with eq_Tg show "integral\<^isup>P M' g = (\\<^isup>+x. f x * g x \M)" unfolding isoton_def by simp qed lemma (in measure_space) positive_integral_null_set: - assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x) = 0" + assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x \M) = 0" proof - - have "(\\<^isup>+ x. u x * indicator N x) = (\\<^isup>+ x. 0)" + have "(\\<^isup>+ x. u x * indicator N x \M) = (\\<^isup>+ x. 0 \M)" proof (intro positive_integral_cong_AE AE_I) show "{x \ space M. u x * indicator N x \ 0} \ N" by (auto simp: indicator_def) @@ -1396,20 +1416,20 @@ lemma (in measure_space) positive_integral_Markov_inequality: assumes borel: "u \ borel_measurable M" and "A \ sets M" and c: "c \ \" - shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x)" + shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" (is "\ ?A \ _ * ?PI") proof - have "?A \ sets M" using `A \ sets M` borel by auto - hence "\ ?A = (\\<^isup>+ x. indicator ?A x)" + hence "\ ?A = (\\<^isup>+ x. indicator ?A x \M)" using positive_integral_indicator by simp - also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x))" + also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x) \M)" proof (rule positive_integral_mono) fix x assume "x \ space M" show "indicator ?A x \ c * (u x * indicator A x)" by (cases "x \ ?A") auto qed - also have "\ = c * (\\<^isup>+ x. u x * indicator A x)" + also have "\ = c * (\\<^isup>+ x. u x * indicator A x \M)" using assms by (auto intro!: positive_integral_cmult borel_measurable_indicator) finally show ?thesis . @@ -1417,11 +1437,11 @@ lemma (in measure_space) positive_integral_0_iff: assumes borel: "u \ borel_measurable M" - shows "positive_integral u = 0 \ \ {x\space M. u x \ 0} = 0" + shows "integral\<^isup>P M u = 0 \ \ {x\space M. u x \ 0} = 0" (is "_ \ \ ?A = 0") proof - have A: "?A \ sets M" using borel by auto - have u: "(\\<^isup>+ x. u x * indicator ?A x) = positive_integral u" + have u: "(\\<^isup>+ x. u x * indicator ?A x \M) = integral\<^isup>P M u" by (auto intro!: positive_integral_cong simp: indicator_def) show ?thesis @@ -1429,10 +1449,10 @@ assume "\ ?A = 0" hence "?A \ null_sets" using `?A \ sets M` by auto from positive_integral_null_set[OF this] - have "0 = (\\<^isup>+ x. u x * indicator ?A x)" by simp - thus "positive_integral u = 0" unfolding u by simp + have "0 = (\\<^isup>+ x. u x * indicator ?A x \M)" by simp + thus "integral\<^isup>P M u = 0" unfolding u by simp next - assume *: "positive_integral u = 0" + assume *: "integral\<^isup>P M u = 0" let "?M n" = "{x \ space M. 1 \ of_nat n * u x}" have "0 = (SUP n. \ (?M n \ ?A))" proof - @@ -1469,34 +1489,34 @@ lemma (in measure_space) positive_integral_restricted: assumes "A \ sets M" - shows "measure_space.positive_integral (restricted_space A) \ f = (\\<^isup>+ x. f x * indicator A x)" - (is "measure_space.positive_integral ?R \ f = positive_integral ?f") + shows "integral\<^isup>P (restricted_space A) f = (\\<^isup>+ x. f x * indicator A x \M)" + (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f") proof - - have msR: "measure_space ?R \" by (rule restricted_measure_space[OF `A \ sets M`]) - then interpret R: measure_space ?R \ . + have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \ sets M`]) + then interpret R: measure_space ?R . have saR: "sigma_algebra ?R" by fact - have *: "R.positive_integral f = R.positive_integral ?f" + have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f" by (intro R.positive_integral_cong) auto show ?thesis - unfolding * R.positive_integral_def positive_integral_def + unfolding * positive_integral_def unfolding simple_function_restricted[OF `A \ sets M`] apply (simp add: SUPR_def) apply (rule arg_cong[where f=Sup]) proof (auto simp add: image_iff simple_integral_restricted[OF `A \ sets M`]) - fix g assume "simple_function (\x. g x * indicator A x)" + fix g assume "simple_function M (\x. g x * indicator A x)" "g \ f" - then show "\x. simple_function x \ x \ (\x. f x * indicator A x) \ - (\\<^isup>Sx. g x * indicator A x) = simple_integral x" + then show "\x. simple_function M x \ x \ (\x. f x * indicator A x) \ + (\\<^isup>Sx. g x * indicator A x \M) = integral\<^isup>S M x" apply (rule_tac exI[of _ "\x. g x * indicator A x"]) by (auto simp: indicator_def le_fun_def) next - fix g assume g: "simple_function g" "g \ (\x. f x * indicator A x)" + fix g assume g: "simple_function M g" "g \ (\x. f x * indicator A x)" then have *: "(\x. g x * indicator A x) = g" "\x. g x * indicator A x = g x" "\x. g x \ f x" by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm) - from g show "\x. simple_function (\xa. x xa * indicator A xa) \ x \ f \ - simple_integral g = simple_integral (\xa. x xa * indicator A xa)" + from g show "\x. simple_function M (\xa. x xa * indicator A xa) \ x \ f \ + integral\<^isup>S M g = integral\<^isup>S M (\xa. x xa * indicator A xa)" using `A \ sets M`[THEN sets_into_space] apply (rule_tac exI[of _ "\x. g x * indicator A x"]) by (fastsimp simp: le_fun_def *) @@ -1505,103 +1525,113 @@ lemma (in measure_space) positive_integral_subalgebra: assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" and sa: "sigma_algebra N" - shows "measure_space.positive_integral N \ f = positive_integral f" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" + and sa: "sigma_algebra N" + shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - interpret N: measure_space N \ using measure_space_subalgebra[OF sa N] . + interpret N: measure_space N using measure_space_subalgebra[OF sa N] . from N.borel_measurable_implies_simple_function_sequence[OF borel] - obtain fs where Nsf: "\i. N.simple_function (fs i)" and "fs \ f" by blast - then have sf: "\i. simple_function (fs i)" - using simple_function_subalgebra[OF _ N sa] by blast - from positive_integral_isoton_simple[OF `fs \ f` sf] N.positive_integral_isoton_simple[OF `fs \ f` Nsf] - show ?thesis unfolding isoton_def simple_integral_def N.simple_integral_def `space N = space M` by simp + obtain fs where Nsf: "\i. simple_function N (fs i)" and "fs \ f" by blast + then have sf: "\i. simple_function M (fs i)" + using simple_function_subalgebra[OF _ N(1,2)] by blast + from N.positive_integral_isoton_simple[OF `fs \ f` Nsf] + have "integral\<^isup>P N f = (SUP i. \x\fs i ` space M. x * N.\ (fs i -` {x} \ space M))" + unfolding isoton_def simple_integral_def `space N = space M` by simp + also have "\ = (SUP i. \x\fs i ` space M. x * \ (fs i -` {x} \ space M))" + using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto + also have "\ = integral\<^isup>P M f" + using positive_integral_isoton_simple[OF `fs \ f` sf] + unfolding isoton_def simple_integral_def `space N = space M` by simp + finally show ?thesis . qed section "Lebesgue Integral" -definition (in measure_space) integrable where - "integrable f \ f \ borel_measurable M \ - (\\<^isup>+ x. Real (f x)) \ \ \ - (\\<^isup>+ x. Real (- f x)) \ \" +definition integrable where + "integrable M f \ f \ borel_measurable M \ + (\\<^isup>+ x. Real (f x) \M) \ \ \ + (\\<^isup>+ x. Real (- f x) \M) \ \" -lemma (in measure_space) integrableD[dest]: - assumes "integrable f" - shows "f \ borel_measurable M" - "(\\<^isup>+ x. Real (f x)) \ \" - "(\\<^isup>+ x. Real (- f x)) \ \" +lemma integrableD[dest]: + assumes "integrable M f" + shows "f \ borel_measurable M" "(\\<^isup>+ x. Real (f x) \M) \ \" "(\\<^isup>+ x. Real (- f x) \M) \ \" using assms unfolding integrable_def by auto -definition (in measure_space) integral (binder "\ " 10) where - "integral f = real ((\\<^isup>+ x. Real (f x))) - real ((\\<^isup>+ x. Real (- f x)))" +definition lebesgue_integral_def: + "integral\<^isup>L M f = real ((\\<^isup>+ x. Real (f x) \M)) - real ((\\<^isup>+ x. Real (- f x) \M))" + +syntax + "_lebesgue_integral" :: "'a \ ('a \ real) \ ('a, 'b) measure_space_scheme \ real" ("\ _. _ \_" [60,61] 110) + +translations + "\ x. f \M" == "CONST integral\<^isup>L M (%x. f)" lemma (in measure_space) integral_cong: - assumes cong: "\x. x \ space M \ f x = g x" - shows "integral f = integral g" - using assms by (simp cong: positive_integral_cong add: integral_def) + assumes "\x. x \ space M \ f x = g x" + shows "integral\<^isup>L M f = integral\<^isup>L M g" + using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def) lemma (in measure_space) integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" - shows "measure_space.integral M \ f = integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "integral\<^isup>L N f = integral\<^isup>L M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact + interpret v: measure_space N + by (rule measure_space_cong) fact+ show ?thesis - unfolding integral_def v.integral_def - by (simp add: positive_integral_cong_measure[OF assms]) + by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def) qed lemma (in measure_space) integral_cong_AE: assumes cong: "AE x. f x = g x" - shows "integral f = integral g" + shows "integral\<^isup>L M f = integral\<^isup>L M g" proof - have "AE x. Real (f x) = Real (g x)" using cong by (rule AE_mp) simp moreover have "AE x. Real (- f x) = Real (- g x)" using cong by (rule AE_mp) simp ultimately show ?thesis - by (simp cong: positive_integral_cong_AE add: integral_def) + by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def) qed lemma (in measure_space) integrable_cong: - "(\x. x \ space M \ f x = g x) \ integrable f \ integrable g" + "(\x. x \ space M \ f x = g x) \ integrable M f \ integrable M g" by (simp cong: positive_integral_cong measurable_cong add: integrable_def) lemma (in measure_space) integral_eq_positive_integral: assumes "\x. 0 \ f x" - shows "integral f = real ((\\<^isup>+ x. Real (f x)))" + shows "integral\<^isup>L M f = real (\\<^isup>+ x. Real (f x) \M)" proof - have "\x. Real (- f x) = 0" using assms by simp - thus ?thesis by (simp del: Real_eq_0 add: integral_def) + thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def) qed lemma (in measure_space) integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - assumes f: "measure_space.integrable M' (\A. \ (T -` A \ space M)) f" - shows "integrable (\x. f (T x))" (is ?P) - and "measure_space.integral M' (\A. \ (T -` A \ space M)) f = (\x. f (T x))" (is ?I) + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + assumes f: "integrable M' f" + shows "integrable M (\x. f (T x))" (is ?P) + and "integral\<^isup>L M' f = (\x. f (T x) \M)" (is ?I) proof - - interpret T: measure_space M' "\A. \ (T -` A \ space M)" - using T by (rule measure_space_vimage) auto + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) from measurable_comp[OF T(2), of f borel] have borel: "(\x. Real (f x)) \ borel_measurable M'" "(\x. Real (- f x)) \ borel_measurable M'" and "(\x. f (T x)) \ borel_measurable M" - using f unfolding T.integrable_def by (auto simp: comp_def) + using f unfolding integrable_def by (auto simp: comp_def) then show ?P ?I - using f unfolding T.integral_def integral_def T.integrable_def integrable_def - unfolding borel[THEN positive_integral_vimage[OF T]] by auto + using f unfolding lebesgue_integral_def integrable_def + by (auto simp: borel[THEN positive_integral_vimage[OF T], OF \]) qed lemma (in measure_space) integral_minus[intro, simp]: - assumes "integrable f" - shows "integrable (\x. - f x)" "(\x. - f x) = - integral f" - using assms by (auto simp: integrable_def integral_def) + assumes "integrable M f" + shows "integrable M (\x. - f x)" "(\x. - f x \M) = - integral\<^isup>L M f" + using assms by (auto simp: integrable_def lebesgue_integral_def) lemma (in measure_space) integral_of_positive_diff: - assumes integrable: "integrable u" "integrable v" + assumes integrable: "integrable M u" "integrable M v" and f_def: "\x. f x = u x - v x" and pos: "\x. 0 \ u x" "\x. 0 \ v x" - shows "integrable f" and "integral f = integral u - integral v" + shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" proof - - let ?PI = positive_integral let "?f x" = "Real (f x)" let "?mf x" = "Real (- f x)" let "?u x" = "Real (u x)" @@ -1615,38 +1645,39 @@ "f \ borel_measurable M" by (auto simp: f_def[symmetric] integrable_def) - have "?PI (\x. Real (u x - v x)) \ ?PI ?u" + have "(\\<^isup>+ x. Real (u x - v x) \M) \ integral\<^isup>P M ?u" using pos by (auto intro!: positive_integral_mono) - moreover have "?PI (\x. Real (v x - u x)) \ ?PI ?v" + moreover have "(\\<^isup>+ x. Real (v x - u x) \M) \ integral\<^isup>P M ?v" using pos by (auto intro!: positive_integral_mono) - ultimately show f: "integrable f" - using `integrable u` `integrable v` `f \ borel_measurable M` + ultimately show f: "integrable M f" + using `integrable M u` `integrable M v` `f \ borel_measurable M` by (auto simp: integrable_def f_def) - hence mf: "integrable (\x. - f x)" .. + hence mf: "integrable M (\x. - f x)" .. have *: "\x. Real (- v x) = 0" "\x. Real (- u x) = 0" using pos by auto have "\x. ?u x + ?mf x = ?v x + ?f x" unfolding f_def using pos by simp - hence "?PI (\x. ?u x + ?mf x) = ?PI (\x. ?v x + ?f x)" by simp - hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)" + hence "(\\<^isup>+ x. ?u x + ?mf x \M) = (\\<^isup>+ x. ?v x + ?f x \M)" by simp + hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) = + real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)" using positive_integral_add[OF u_borel mf_borel] using positive_integral_add[OF v_borel f_borel] by auto - then show "integral f = integral u - integral v" - using f mf `integrable u` `integrable v` - unfolding integral_def integrable_def * - by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u") + then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" + using f mf `integrable M u` `integrable M v` + unfolding lebesgue_integral_def integrable_def * + by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u") (auto simp add: field_simps) qed lemma (in measure_space) integral_linear: - assumes "integrable f" "integrable g" and "0 \ a" - shows "integrable (\t. a * f t + g t)" - and "(\ t. a * f t + g t) = a * integral f + integral g" + assumes "integrable M f" "integrable M g" and "0 \ a" + shows "integrable M (\t. a * f t + g t)" + and "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" proof - - let ?PI = positive_integral + let ?PI = "integral\<^isup>P M" let "?f x" = "Real (f x)" let "?g x" = "Real (g x)" let "?mf x" = "Real (- f x)" @@ -1670,37 +1701,36 @@ positive_integral_linear[OF pos] positive_integral_linear[OF neg] - have "integrable ?p" "integrable ?n" + have "integrable M ?p" "integrable M ?n" "\t. a * f t + g t = ?p t - ?n t" "\t. 0 \ ?p t" "\t. 0 \ ?n t" using assms p n unfolding integrable_def * linear by auto note diff = integral_of_positive_diff[OF this] - show "integrable (\t. a * f t + g t)" by (rule diff) + show "integrable M (\t. a * f t + g t)" by (rule diff) - from assms show "(\ t. a * f t + g t) = a * integral f + integral g" - unfolding diff(2) unfolding integral_def * linear integrable_def + from assms show "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" + unfolding diff(2) unfolding lebesgue_integral_def * linear integrable_def by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg") (auto simp add: field_simps zero_le_mult_iff) qed lemma (in measure_space) integral_add[simp, intro]: - assumes "integrable f" "integrable g" - shows "integrable (\t. f t + g t)" - and "(\ t. f t + g t) = integral f + integral g" + assumes "integrable M f" "integrable M g" + shows "integrable M (\t. f t + g t)" + and "(\ t. f t + g t \M) = integral\<^isup>L M f + integral\<^isup>L M g" using assms integral_linear[where a=1] by auto lemma (in measure_space) integral_zero[simp, intro]: - shows "integrable (\x. 0)" - and "(\ x.0) = 0" - unfolding integrable_def integral_def + shows "integrable M (\x. 0)" "(\ x.0 \M) = 0" + unfolding integrable_def lebesgue_integral_def by (auto simp add: borel_measurable_const) lemma (in measure_space) integral_cmult[simp, intro]: - assumes "integrable f" - shows "integrable (\t. a * f t)" (is ?P) - and "(\ t. a * f t) = a * integral f" (is ?I) + assumes "integrable M f" + shows "integrable M (\t. a * f t)" (is ?P) + and "(\ t. a * f t \M) = a * integral\<^isup>L M f" (is ?I) proof - - have "integrable (\t. a * f t) \ (\ t. a * f t) = a * integral f" + have "integrable M (\t. a * f t) \ (\ t. a * f t \M) = a * integral\<^isup>L M f" proof (cases rule: le_cases) assume "0 \ a" show ?thesis using integral_linear[OF assms integral_zero(1) `0 \ a`] @@ -1716,56 +1746,56 @@ qed lemma (in measure_space) integral_multc: - assumes "integrable f" - shows "(\ x. f x * c) = integral f * c" + assumes "integrable M f" + shows "(\ x. f x * c \M) = integral\<^isup>L M f * c" unfolding mult_commute[of _ c] integral_cmult[OF assms] .. lemma (in measure_space) integral_mono_AE: - assumes fg: "integrable f" "integrable g" + assumes fg: "integrable M f" "integrable M g" and mono: "AE t. f t \ g t" - shows "integral f \ integral g" + shows "integral\<^isup>L M f \ integral\<^isup>L M g" proof - have "AE x. Real (f x) \ Real (g x)" using mono by (rule AE_mp) (auto intro!: AE_cong) - moreover have "AE x. Real (- g x) \ Real (- f x)" + moreover have "AE x. Real (- g x) \ Real (- f x)" using mono by (rule AE_mp) (auto intro!: AE_cong) ultimately show ?thesis using fg - by (auto simp: integral_def integrable_def diff_minus + by (auto simp: lebesgue_integral_def integrable_def diff_minus intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE) qed lemma (in measure_space) integral_mono: - assumes fg: "integrable f" "integrable g" + assumes fg: "integrable M f" "integrable M g" and mono: "\t. t \ space M \ f t \ g t" - shows "integral f \ integral g" + shows "integral\<^isup>L M f \ integral\<^isup>L M g" apply (rule integral_mono_AE[OF fg]) using mono by (rule AE_cong) auto lemma (in measure_space) integral_diff[simp, intro]: - assumes f: "integrable f" and g: "integrable g" - shows "integrable (\t. f t - g t)" - and "(\ t. f t - g t) = integral f - integral g" + assumes f: "integrable M f" and g: "integrable M g" + shows "integrable M (\t. f t - g t)" + and "(\ t. f t - g t \M) = integral\<^isup>L M f - integral\<^isup>L M g" using integral_add[OF f integral_minus(1)[OF g]] unfolding diff_minus integral_minus(2)[OF g] by auto lemma (in measure_space) integral_indicator[simp, intro]: assumes "a \ sets M" and "\ a \ \" - shows "integral (indicator a) = real (\ a)" (is ?int) - and "integrable (indicator a)" (is ?able) + shows "integral\<^isup>L M (indicator a) = real (\ a)" (is ?int) + and "integrable M (indicator a)" (is ?able) proof - have *: "\A x. Real (indicator A x) = indicator A x" "\A x. Real (- indicator A x) = 0" unfolding indicator_def by auto show ?int ?able - using assms unfolding integral_def integrable_def + using assms unfolding lebesgue_integral_def integrable_def by (auto simp: * positive_integral_indicator borel_measurable_indicator) qed lemma (in measure_space) integral_cmul_indicator: assumes "A \ sets M" and "c \ 0 \ \ A \ \" - shows "integrable (\x. c * indicator A x)" (is ?P) - and "(\x. c * indicator A x) = c * real (\ A)" (is ?I) + shows "integrable M (\x. c * indicator A x)" (is ?P) + and "(\x. c * indicator A x \M) = c * real (\ A)" (is ?I) proof - show ?P proof (cases "c = 0") @@ -1779,9 +1809,9 @@ qed lemma (in measure_space) integral_setsum[simp, intro]: - assumes "\n. n \ S \ integrable (f n)" - shows "(\x. \ i \ S. f i x) = (\ i \ S. integral (f i))" (is "?int S") - and "integrable (\x. \ i \ S. f i x)" (is "?I S") + assumes "\n. n \ S \ integrable M (f n)" + shows "(\x. (\ i \ S. f i x) \M) = (\ i \ S. integral\<^isup>L M (f i))" (is "?int S") + and "integrable M (\x. \ i \ S. f i x)" (is "?I S") proof - have "?int S \ ?I S" proof (cases "finite S") @@ -1792,8 +1822,8 @@ qed lemma (in measure_space) integrable_abs: - assumes "integrable f" - shows "integrable (\ x. \f x\)" + assumes "integrable M f" + shows "integrable M (\ x. \f x\)" proof - have *: "\x. Real \f x\ = Real (f x) + Real (- f x)" @@ -1808,56 +1838,55 @@ lemma (in measure_space) integral_subalgebra: assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" and sa: "sigma_algebra N" - shows "measure_space.integrable N \ f \ integrable f" (is ?P) - and "measure_space.integral N \ f = integral f" (is ?I) + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" and sa: "sigma_algebra N" + shows "integrable N f \ integrable M f" (is ?P) + and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I) proof - - interpret N: measure_space N \ using measure_space_subalgebra[OF sa N] . + interpret N: measure_space N using measure_space_subalgebra[OF sa N] . have "(\x. Real (f x)) \ borel_measurable N" "(\x. Real (- f x)) \ borel_measurable N" using borel by auto note * = this[THEN positive_integral_subalgebra[OF _ N sa]] have "f \ borel_measurable M \ f \ borel_measurable N" using assms unfolding measurable_def by auto - then show ?P ?I unfolding integrable_def N.integrable_def integral_def N.integral_def - unfolding * by auto + then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def) qed lemma (in measure_space) integrable_bound: - assumes "integrable f" + assumes "integrable M f" and f: "\x. x \ space M \ 0 \ f x" "\x. x \ space M \ \g x\ \ f x" assumes borel: "g \ borel_measurable M" - shows "integrable g" + shows "integrable M g" proof - - have "(\\<^isup>+ x. Real (g x)) \ (\\<^isup>+ x. Real \g x\)" + have "(\\<^isup>+ x. Real (g x) \M) \ (\\<^isup>+ x. Real \g x\ \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x) \M)" using f by (auto intro!: positive_integral_mono) also have "\ < \" - using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have pos: "(\\<^isup>+ x. Real (g x)) < \" . + using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) + finally have pos: "(\\<^isup>+ x. Real (g x) \M) < \" . - have "(\\<^isup>+ x. Real (- g x)) \ (\\<^isup>+ x. Real (\g x\))" + have "(\\<^isup>+ x. Real (- g x) \M) \ (\\<^isup>+ x. Real (\g x\) \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x) \M)" using f by (auto intro!: positive_integral_mono) also have "\ < \" - using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have neg: "(\\<^isup>+ x. Real (- g x)) < \" . + using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) + finally have neg: "(\\<^isup>+ x. Real (- g x) \M) < \" . from neg pos borel show ?thesis unfolding integrable_def by auto qed lemma (in measure_space) integrable_abs_iff: - "f \ borel_measurable M \ integrable (\ x. \f x\) \ integrable f" + "f \ borel_measurable M \ integrable M (\ x. \f x\) \ integrable M f" by (auto intro!: integrable_bound[where g=f] integrable_abs) lemma (in measure_space) integrable_max: - assumes int: "integrable f" "integrable g" - shows "integrable (\ x. max (f x) (g x))" + assumes int: "integrable M f" "integrable M g" + shows "integrable M (\ x. max (f x) (g x))" proof (rule integrable_bound) - show "integrable (\x. \f x\ + \g x\)" + show "integrable M (\x. \f x\ + \g x\)" using int by (simp add: integrable_abs) show "(\x. max (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto @@ -1868,10 +1897,10 @@ qed lemma (in measure_space) integrable_min: - assumes int: "integrable f" "integrable g" - shows "integrable (\ x. min (f x) (g x))" + assumes int: "integrable M f" "integrable M g" + shows "integrable M (\ x. min (f x) (g x))" proof (rule integrable_bound) - show "integrable (\x. \f x\ + \g x\)" + show "integrable M (\x. \f x\ + \g x\)" using int by (simp add: integrable_abs) show "(\x. min (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto @@ -1882,33 +1911,33 @@ qed lemma (in measure_space) integral_triangle_inequality: - assumes "integrable f" - shows "\integral f\ \ (\x. \f x\)" + assumes "integrable M f" + shows "\integral\<^isup>L M f\ \ (\x. \f x\ \M)" proof - - have "\integral f\ = max (integral f) (- integral f)" by auto - also have "\ \ (\x. \f x\)" + have "\integral\<^isup>L M f\ = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto + also have "\ \ (\x. \f x\ \M)" using assms integral_minus(2)[of f, symmetric] by (auto intro!: integral_mono integrable_abs simp del: integral_minus) finally show ?thesis . qed lemma (in measure_space) integral_positive: - assumes "integrable f" "\x. x \ space M \ 0 \ f x" - shows "0 \ integral f" + assumes "integrable M f" "\x. x \ space M \ 0 \ f x" + shows "0 \ integral\<^isup>L M f" proof - - have "0 = (\x. 0)" by (auto simp: integral_zero) - also have "\ \ integral f" + have "0 = (\x. 0 \M)" by (auto simp: integral_zero) + also have "\ \ integral\<^isup>L M f" using assms by (rule integral_mono[OF integral_zero(1)]) finally show ?thesis . qed lemma (in measure_space) integral_monotone_convergence_pos: - assumes i: "\i. integrable (f i)" and mono: "\x. mono (\n. f n x)" + assumes i: "\i. integrable M (f i)" and mono: "\x. mono (\n. f n x)" and pos: "\x i. 0 \ f i x" and lim: "\x. (\i. f i x) ----> u x" - and ilim: "(\i. integral (f i)) ----> x" - shows "integrable u" - and "integral u = x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + shows "integrable M u" + and "integral\<^isup>L M u = x" proof - { fix x have "0 \ u x" using mono pos[of 0 x] incseq_le[OF _ lim, of x 0] @@ -1928,43 +1957,42 @@ hence borel_u: "u \ borel_measurable M" using pos_u by (auto simp: borel_measurable_Real_eq SUP_F) - have integral_eq: "\n. (\\<^isup>+ x. Real (f n x)) = Real (integral (f n))" - using i unfolding integral_def integrable_def by (auto simp: Real_real) + have integral_eq: "\n. (\\<^isup>+ x. Real (f n x) \M) = Real (integral\<^isup>L M (f n))" + using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real) - have pos_integral: "\n. 0 \ integral (f n)" + have pos_integral: "\n. 0 \ integral\<^isup>L M (f n)" using pos i by (auto simp: integral_positive) hence "0 \ x" using LIMSEQ_le_const[OF ilim, of 0] by auto - have "(\i. (\\<^isup>+ x. Real (f i x))) \ (\\<^isup>+ x. Real (u x))" + have "(\i. (\\<^isup>+ x. Real (f i x) \M)) \ (\\<^isup>+ x. Real (u x) \M)" proof (rule positive_integral_isoton) from SUP_F mono pos show "(\i x. Real (f i x)) \ (\x. Real (u x))" unfolding isoton_fun_expand by (auto simp: isoton_def mono_def) qed (rule borel_f) - hence pI: "(\\<^isup>+ x. Real (u x)) = - (SUP n. (\\<^isup>+ x. Real (f n x)))" + hence pI: "(\\<^isup>+ x. Real (u x) \M) = (SUP n. (\\<^isup>+ x. Real (f n x) \M))" unfolding isoton_def by simp also have "\ = Real x" unfolding integral_eq proof (rule SUP_eq_LIMSEQ[THEN iffD2]) - show "mono (\n. integral (f n))" + show "mono (\n. integral\<^isup>L M (f n))" using mono i by (auto simp: mono_def intro!: integral_mono) - show "\n. 0 \ integral (f n)" using pos_integral . + show "\n. 0 \ integral\<^isup>L M (f n)" using pos_integral . show "0 \ x" using `0 \ x` . - show "(\n. integral (f n)) ----> x" using ilim . + show "(\n. integral\<^isup>L M (f n)) ----> x" using ilim . qed - finally show "integrable u" "integral u = x" using borel_u `0 \ x` - unfolding integrable_def integral_def by auto + finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \ x` + unfolding integrable_def lebesgue_integral_def by auto qed lemma (in measure_space) integral_monotone_convergence: - assumes f: "\i. integrable (f i)" and "mono f" + assumes f: "\i. integrable M (f i)" and "mono f" and lim: "\x. (\i. f i x) ----> u x" - and ilim: "(\i. integral (f i)) ----> x" - shows "integrable u" - and "integral u = x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + shows "integrable M u" + and "integral\<^isup>L M u = x" proof - - have 1: "\i. integrable (\x. f i x - f 0 x)" + have 1: "\i. integrable M (\x. f i x - f 0 x)" using f by (auto intro!: integral_diff) have 2: "\x. mono (\n. f n x - f 0 x)" using `mono f` unfolding mono_def le_fun_def by auto @@ -1972,43 +2000,43 @@ unfolding mono_def le_fun_def by (auto simp: field_simps) have 4: "\x. (\i. f i x - f 0 x) ----> u x - f 0 x" using lim by (auto intro!: LIMSEQ_diff) - have 5: "(\i. (\x. f i x - f 0 x)) ----> x - integral (f 0)" + have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^isup>L M (f 0)" using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff) note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5] - have "integrable (\x. (u x - f 0 x) + f 0 x)" + have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integral_add(1)) - with diff(2) f show "integrable u" "integral u = x" + with diff(2) f show "integrable M u" "integral\<^isup>L M u = x" by (auto simp: integral_diff) qed lemma (in measure_space) integral_0_iff: - assumes "integrable f" - shows "(\x. \f x\) = 0 \ \ {x\space M. f x \ 0} = 0" + assumes "integrable M f" + shows "(\x. \f x\ \M) = 0 \ \ {x\space M. f x \ 0} = 0" proof - have *: "\x. Real (- \f x\) = 0" by auto - have "integrable (\x. \f x\)" using assms by (rule integrable_abs) + have "integrable M (\x. \f x\)" using assms by (rule integrable_abs) hence "(\x. Real (\f x\)) \ borel_measurable M" - "(\\<^isup>+ x. Real \f x\) \ \" unfolding integrable_def by auto + "(\\<^isup>+ x. Real \f x\ \M) \ \" unfolding integrable_def by auto from positive_integral_0_iff[OF this(1)] this(2) - show ?thesis unfolding integral_def * + show ?thesis unfolding lebesgue_integral_def * by (simp add: real_of_pextreal_eq_0) qed lemma (in measure_space) positive_integral_omega: assumes "f \ borel_measurable M" - and "positive_integral f \ \" + and "integral\<^isup>P M f \ \" shows "\ (f -` {\} \ space M) = 0" proof - - have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x)" + have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \ \] by simp - also have "\ \ positive_integral f" + also have "\ \ integral\<^isup>P M f" by (auto intro!: positive_integral_mono simp: indicator_def) finally show ?thesis using assms(2) by (cases ?thesis) auto qed lemma (in measure_space) positive_integral_omega_AE: - assumes "f \ borel_measurable M" "positive_integral f \ \" shows "AE x. f x \ \" + assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x. f x \ \" proof (rule AE_I) show "\ (f -` {\} \ space M) = 0" by (rule positive_integral_omega[OF assms]) @@ -2017,39 +2045,39 @@ qed auto lemma (in measure_space) simple_integral_omega: - assumes "simple_function f" - and "simple_integral f \ \" + assumes "simple_function M f" + and "integral\<^isup>S M f \ \" shows "\ (f -` {\} \ space M) = 0" proof (rule positive_integral_omega) show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) - show "positive_integral f \ \" + show "integral\<^isup>P M f \ \" using assms by (simp add: positive_integral_eq_simple_integral) qed lemma (in measure_space) integral_real: fixes f :: "'a \ pextreal" assumes "AE x. f x \ \" - shows "(\x. real (f x)) = real (positive_integral f)" (is ?plus) - and "(\x. - real (f x)) = - real (positive_integral f)" (is ?minus) + shows "(\x. real (f x) \M) = real (integral\<^isup>P M f)" (is ?plus) + and "(\x. - real (f x) \M) = - real (integral\<^isup>P M f)" (is ?minus) proof - - have "(\\<^isup>+ x. Real (real (f x))) = positive_integral f" + have "(\\<^isup>+ x. Real (real (f x)) \M) = integral\<^isup>P M f" apply (rule positive_integral_cong_AE) apply (rule AE_mp[OF assms(1)]) by (auto intro!: AE_cong simp: Real_real) moreover - have "(\\<^isup>+ x. Real (- real (f x))) = (\\<^isup>+ x. 0)" + have "(\\<^isup>+ x. Real (- real (f x)) \M) = (\\<^isup>+ x. 0 \M)" by (intro positive_integral_cong) auto ultimately show ?plus ?minus - by (auto simp: integral_def integrable_def) + by (auto simp: lebesgue_integral_def integrable_def) qed lemma (in measure_space) integral_dominated_convergence: - assumes u: "\i. integrable (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" - and w: "integrable w" "\x. x \ space M \ 0 \ w x" + assumes u: "\i. integrable M (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" + and w: "integrable M w" "\x. x \ space M \ 0 \ w x" and u': "\x. x \ space M \ (\i. u i x) ----> u' x" - shows "integrable u'" - and "(\i. (\x. \u i x - u' x\)) ----> 0" (is "?lim_diff") - and "(\i. integral (u i)) ----> integral u'" (is ?lim) + shows "integrable M u'" + and "(\i. (\x. \u i x - u' x\ \M)) ----> 0" (is "?lim_diff") + and "(\i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim) proof - { fix x j assume x: "x \ space M" from u'[OF x] have "(\i. \u i x\) ----> \u' x\" by (rule LIMSEQ_imp_rabs) @@ -2061,9 +2089,9 @@ have u'_borel: "u' \ borel_measurable M" using u' by (blast intro: borel_measurable_LIMSEQ[of u]) - show "integrable u'" + show "integrable M u'" proof (rule integrable_bound) - show "integrable w" by fact + show "integrable M w" by fact show "u' \ borel_measurable M" by fact next fix x assume x: "x \ space M" @@ -2072,8 +2100,8 @@ qed let "?diff n x" = "2 * w x - \u n x - u' x\" - have diff: "\n. integrable (\x. \u n x - u' x\)" - using w u `integrable u'` + have diff: "\n. integrable M (\x. \u n x - u' x\)" + using w u `integrable M u'` by (auto intro!: integral_add integral_diff integral_cmult integrable_abs) { fix j x assume x: "x \ space M" @@ -2083,31 +2111,31 @@ finally have "\u j x - u' x\ \ 2 * w x" by simp } note diff_less_2w = this - have PI_diff: "\m n. (\\<^isup>+ x. Real (?diff (m + n) x)) = - (\\<^isup>+ x. Real (2 * w x)) - (\\<^isup>+ x. Real \u (m + n) x - u' x\)" + have PI_diff: "\m n. (\\<^isup>+ x. Real (?diff (m + n) x) \M) = + (\\<^isup>+ x. Real (2 * w x) \M) - (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" using diff w diff_less_2w by (subst positive_integral_diff[symmetric]) (auto simp: integrable_def intro!: positive_integral_cong) - have "integrable (\x. 2 * w x)" + have "integrable M (\x. 2 * w x)" using w by (auto intro: integral_cmult) - hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x)) \ \" and + hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x) \M) \ \" and borel_2w: "(\x. Real (2 * w x)) \ borel_measurable M" unfolding integrable_def by auto - have "(INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\)) = 0" (is "?lim_SUP = 0") + have "(INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)) = 0" (is "?lim_SUP = 0") proof cases - assume eq_0: "(\\<^isup>+ x. Real (2 * w x)) = 0" - have "\i. (\\<^isup>+ x. Real \u i x - u' x\) \ (\\<^isup>+ x. Real (2 * w x))" + assume eq_0: "(\\<^isup>+ x. Real (2 * w x) \M) = 0" + have "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) \ (\\<^isup>+ x. Real (2 * w x) \M)" proof (rule positive_integral_mono) fix i x assume "x \ space M" from diff_less_2w[OF this, of i] show "Real \u i x - u' x\ \ Real (2 * w x)" by auto qed - hence "\i. (\\<^isup>+ x. Real \u i x - u' x\) = 0" using eq_0 by auto + hence "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) = 0" using eq_0 by auto thus ?thesis by simp next - assume neq_0: "(\\<^isup>+ x. Real (2 * w x)) \ 0" - have "(\\<^isup>+ x. Real (2 * w x)) = (\\<^isup>+ x. SUP n. INF m. Real (?diff (m + n) x))" + assume neq_0: "(\\<^isup>+ x. Real (2 * w x) \M) \ 0" + have "(\\<^isup>+ x. Real (2 * w x) \M) = (\\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \M)" proof (rule positive_integral_cong, subst add_commute) fix x assume x: "x \ space M" show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))" @@ -2119,22 +2147,22 @@ thus "(\i. ?diff i x) ----> 2 * w x" by simp qed qed - also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x)))" + also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x) \M))" using u'_borel w u unfolding integrable_def by (auto intro!: positive_integral_lim_INF) - also have "\ = (\\<^isup>+ x. Real (2 * w x)) - - (INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\))" + also have "\ = (\\<^isup>+ x. Real (2 * w x) \M) - + (INF n. SUP m. \\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus .. finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0) qed - + have [simp]: "\n m x. Real (- \u (m + n) x - u' x\) = 0" by auto - have [simp]: "\n m. (\\<^isup>+ x. Real \u (m + n) x - u' x\) = - Real ((\x. \u (n + m) x - u' x\))" - using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real) + have [simp]: "\n m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M) = + Real ((\x. \u (n + m) x - u' x\ \M))" + using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real) - have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\)) \ ?lim_SUP" + have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)) \ ?lim_SUP" (is "?lim_INF \ _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP) hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP) @@ -2143,28 +2171,28 @@ proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from LIMSEQ_D[OF `?lim_diff` this] - obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\) < r" + obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\ \M) < r" using diff by (auto simp: integral_positive) - show "\N. \n\N. norm (integral (u n) - integral u') < r" + show "\N. \n\N. norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" proof (safe intro!: exI[of _ N]) fix n assume "N \ n" - have "\integral (u n) - integral u'\ = \(\x. u n x - u' x)\" - using u `integrable u'` by (auto simp: integral_diff) - also have "\ \ (\x. \u n x - u' x\)" using u `integrable u'` + have "\integral\<^isup>L M (u n) - integral\<^isup>L M u'\ = \(\x. u n x - u' x \M)\" + using u `integrable M u'` by (auto simp: integral_diff) + also have "\ \ (\x. \u n x - u' x\ \M)" using u `integrable M u'` by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff) also note N[OF `N \ n`] - finally show "norm (integral (u n) - integral u') < r" by simp + finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp qed qed qed lemma (in measure_space) integral_sums: - assumes borel: "\i. integrable (f i)" + assumes borel: "\i. integrable M (f i)" and summable: "\x. x \ space M \ summable (\i. \f i x\)" - and sums: "summable (\i. (\x. \f i x\))" - shows "integrable (\x. (\i. f i x))" (is "integrable ?S") - and "(\i. integral (f i)) sums (\x. (\i. f i x))" (is ?integral) + and sums: "summable (\i. (\x. \f i x\ \M))" + shows "integrable M (\x. (\i. f i x))" (is "integrable M ?S") + and "(\i. integral\<^isup>L M (f i)) sums (\x. (\i. f i x) \M)" (is ?integral) proof - have "\x\space M. \w. (\i. \f i x\) sums w" using summable unfolding summable_def by auto @@ -2173,10 +2201,10 @@ let "?w y" = "if y \ space M then w y else 0" - obtain x where abs_sum: "(\i. (\x. \f i x\)) sums x" + obtain x where abs_sum: "(\i. (\x. \f i x\ \M)) sums x" using sums unfolding summable_def .. - have 1: "\n. integrable (\x. \i = 0..n. integrable M (\x. \i = 0.. space M" @@ -2185,21 +2213,21 @@ finally have "\\i = 0.. \ ?w x" by simp } note 2 = this - have 3: "integrable ?w" + have 3: "integrable M ?w" proof (rule integral_monotone_convergence(1)) let "?F n y" = "(\i = 0..f i y\)" let "?w' n y" = "if y \ space M then ?F n y else 0" - have "\n. integrable (?F n)" + have "\n. integrable M (?F n)" using borel by (auto intro!: integral_setsum integrable_abs) - thus "\n. integrable (?w' n)" by (simp cong: integrable_cong) + thus "\n. integrable M (?w' n)" by (simp cong: integrable_cong) show "mono ?w'" by (auto simp: mono_def le_fun_def intro!: setsum_mono2) { fix x show "(\n. ?w' n x) ----> ?w x" using w by (cases "x \ space M") (simp_all add: LIMSEQ_const sums_def) } - have *: "\n. integral (?w' n) = (\i = 0..< n. (\x. \f i x\))" + have *: "\n. integral\<^isup>L M (?w' n) = (\i = 0..< n. (\x. \f i x\ \M))" using borel by (simp add: integral_setsum integrable_abs cong: integral_cong) from abs_sum - show "(\i. integral (?w' i)) ----> x" unfolding * sums_def . + show "(\i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def . qed have 4: "\x. x \ space M \ 0 \ ?w x" using 2[of _ 0] by simp @@ -2210,7 +2238,7 @@ note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5] - from int show "integrable ?S" by simp + from int show "integrable M ?S" by simp show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel] using int(2) by simp @@ -2224,12 +2252,12 @@ and enum_zero: "enum ` (-S) \ {0}" and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" and abs_summable: "summable (\r. \enum r * real (\ (f -` {enum r} \ space M))\)" - shows "integrable f" - and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral f" (is ?sums) + shows "integrable M f" + and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral\<^isup>L M f" (is ?sums) proof - let "?A r" = "f -` {enum r} \ space M" let "?F r x" = "enum r * indicator (?A r) x" - have enum_eq: "\r. enum r * real (\ (?A r)) = integral (?F r)" + have enum_eq: "\r. enum r * real (\ (?A r)) = integral\<^isup>L M (?F r)" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) { fix x assume "x \ space M" @@ -2250,19 +2278,19 @@ by (auto intro!: sums_single simp: F F_abs) } note F_sums_f = this(1) and F_abs_sums_f = this(2) - have int_f: "integral f = (\x. \r. ?F r x)" "integrable f = integrable (\x. \r. ?F r x)" + have int_f: "integral\<^isup>L M f = (\x. (\r. ?F r x) \M)" "integrable M f = integrable M (\x. \r. ?F r x)" using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) { fix r - have "(\x. \?F r x\) = (\x. \enum r\ * indicator (?A r) x)" + have "(\x. \?F r x\ \M) = (\x. \enum r\ * indicator (?A r) x \M)" by (auto simp: indicator_def intro!: integral_cong) also have "\ = \enum r\ * real (\ (?A r))" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - finally have "(\x. \?F r x\) = \enum r * real (\ (?A r))\" + finally have "(\x. \?F r x\ \M) = \enum r * real (\ (?A r))\" by (simp add: abs_mult_pos real_pextreal_pos) } note int_abs_F = this - have 1: "\i. integrable (\x. ?F i x)" + have 1: "\i. integrable M (\x. ?F i x)" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) have 2: "\x. x \ space M \ summable (\i. \?F i x\)" @@ -2272,7 +2300,7 @@ show ?sums unfolding enum_eq int_f by simp from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] - show "integrable f" unfolding int_f by simp + show "integrable M f" unfolding int_f by simp qed section "Lebesgue integration on finite space" @@ -2280,8 +2308,8 @@ lemma (in measure_space) integral_on_finite: assumes f: "f \ borel_measurable M" and finite: "finite (f`space M)" and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" - shows "integrable f" - and "(\x. f x) = + shows "integrable M f" + and "(\x. f x \M) = (\ r \ f`space M. r * real (\ (f -` {r} \ space M)))" (is "?integral") proof - let "?A r" = "f -` {r} \ space M" @@ -2295,40 +2323,40 @@ finally have "f x = ?S x" . } note f_eq = this - have f_eq_S: "integrable f \ integrable ?S" "integral f = integral ?S" + have f_eq_S: "integrable M f \ integrable M ?S" "integral\<^isup>L M f = integral\<^isup>L M ?S" by (auto intro!: integrable_cong integral_cong simp only: f_eq) - show "integrable f" ?integral using fin f f_eq_S + show "integrable M f" ?integral using fin f f_eq_S by (simp_all add: integral_cmul_indicator borel_measurable_vimage) qed -lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f" +lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f" unfolding simple_function_def using finite_space by auto lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \ borel_measurable M" by (auto intro: borel_measurable_simple_function) lemma (in finite_measure_space) positive_integral_finite_eq_setsum: - "positive_integral f = (\x \ space M. f x * \ {x})" + "integral\<^isup>P M f = (\x \ space M. f x * \ {x})" proof - - have *: "positive_integral f = (\\<^isup>+ x. \y\space M. f y * indicator {y} x)" + have *: "integral\<^isup>P M f = (\\<^isup>+ x. (\y\space M. f y * indicator {y} x) \M)" by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space]) show ?thesis unfolding * using borel_measurable_finite[of f] by (simp add: positive_integral_setsum positive_integral_cmult_indicator) qed lemma (in finite_measure_space) integral_finite_singleton: - shows "integrable f" - and "integral f = (\x \ space M. f x * real (\ {x}))" (is ?I) + shows "integrable M f" + and "integral\<^isup>L M f = (\x \ space M. f x * real (\ {x}))" (is ?I) proof - have [simp]: - "(\\<^isup>+ x. Real (f x)) = (\x \ space M. Real (f x) * \ {x})" - "(\\<^isup>+ x. Real (- f x)) = (\x \ space M. Real (- f x) * \ {x})" + "(\\<^isup>+ x. Real (f x) \M) = (\x \ space M. Real (f x) * \ {x})" + "(\\<^isup>+ x. Real (- f x) \M) = (\x \ space M. Real (- f x) * \ {x})" unfolding positive_integral_finite_eq_setsum by auto - show "integrable f" using finite_space finite_measure + show "integrable M f" using finite_space finite_measure by (simp add: setsum_\ integrable_def) show ?I using finite_measure - apply (simp add: integral_def real_of_pextreal_setsum[symmetric] + apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric] real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric]) by (rule setsum_cong) (simp_all split: split_if) qed diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Wed Feb 02 12:34:45 2011 +0100 @@ -1,7 +1,7 @@ (* Author: Robert Himmelmann, TU Muenchen *) header {* Lebsegue measure *} theory Lebesgue_Measure - imports Product_Measure Complete_Measure + imports Product_Measure begin subsection {* Standard Cubes *} @@ -42,10 +42,16 @@ by (auto simp add:dist_norm) qed -definition lebesgue :: "'a::ordered_euclidean_space algebra" where - "lebesgue = \ space = UNIV, sets = {A. \n. (indicator A :: 'a \ real) integrable_on cube n} \" +lemma cube_subset_Suc[intro]: "cube n \ cube (Suc n)" + unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto -definition "lmeasure A = (SUP n. Real (integral (cube n) (indicator A)))" +lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" + unfolding Pi_def by auto + +definition lebesgue :: "'a::ordered_euclidean_space measure_space" where + "lebesgue = \ space = UNIV, + sets = {A. \n. (indicator A :: 'a \ real) integrable_on cube n}, + measure = \A. SUP n. Real (integral (cube n) (indicator A)) \" lemma space_lebesgue[simp]: "space lebesgue = UNIV" unfolding lebesgue_def by simp @@ -106,12 +112,12 @@ qed (auto intro: LIMSEQ_indicator_UN simp: cube_def) qed simp -interpretation lebesgue: measure_space lebesgue lmeasure +interpretation lebesgue: measure_space lebesgue proof have *: "indicator {} = (\x. 0 :: real)" by (simp add: fun_eq_iff) - show "lmeasure {} = 0" by (simp add: integral_0 * lmeasure_def) + show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def) next - show "countably_additive lebesgue lmeasure" + show "countably_additive lebesgue (measure lebesgue)" proof (intro countably_additive_def[THEN iffD2] allI impI) fix A :: "nat \ 'b set" assume rA: "range A \ sets lebesgue" "disjoint_family A" then have A[simp, intro]: "\i n. (indicator (A i) :: _ \ real) integrable_on cube n" @@ -122,8 +128,8 @@ assume "(\i. A i) \ sets lebesgue" then have UN_A[simp, intro]: "\i n. (indicator (\i. A i) :: _ \ real) integrable_on cube n" by (auto dest: lebesgueD) - show "(\\<^isub>\n. lmeasure (A n)) = lmeasure (\i. A i)" unfolding lmeasure_def - proof (subst psuminf_SUP_eq) + show "(\\<^isub>\n. measure lebesgue (A n)) = measure lebesgue (\i. A i)" + proof (simp add: lebesgue_def, subst psuminf_SUP_eq) fix n i show "Real (?m n i) \ Real (?m (Suc n) i)" using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le) next @@ -213,20 +219,19 @@ using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI) lemma lmeasure_eq_0: - fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lmeasure S = 0" + fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lebesgue.\ S = 0" proof - have "\n. integral (cube n) (indicator S :: 'a\real) = 0" - unfolding integral_def using assms - by (intro some1_equality ex_ex1I has_integral_unique) - (auto simp: cube_def negligible_def intro: ) - then show ?thesis unfolding lmeasure_def by auto + unfolding lebesgue_integral_def using assms + by (intro integral_unique some1_equality ex_ex1I) + (auto simp: cube_def negligible_def) + then show ?thesis by (auto simp: lebesgue_def) qed lemma lmeasure_iff_LIMSEQ: assumes "A \ sets lebesgue" "0 \ m" - shows "lmeasure A = Real m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" - unfolding lmeasure_def -proof (intro SUP_eq_LIMSEQ) + shows "lebesgue.\ A = Real m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" +proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ) show "mono (\n. integral (cube n) (indicator A::_=>real))" using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD) fix n show "0 \ integral (cube n) (indicator A::_=>real)" @@ -253,7 +258,7 @@ lemma lmeasure_finite_has_integral: fixes s :: "'a::ordered_euclidean_space set" - assumes "s \ sets lebesgue" "lmeasure s = Real m" "0 \ m" + assumes "s \ sets lebesgue" "lebesgue.\ s = Real m" "0 \ m" shows "(indicator s has_integral m) UNIV" proof - let ?I = "indicator :: 'a set \ 'a \ real" @@ -295,9 +300,9 @@ unfolding m by (intro integrable_integral **) qed -lemma lmeasure_finite_integrable: assumes "s \ sets lebesgue" "lmeasure s \ \" +lemma lmeasure_finite_integrable: assumes "s \ sets lebesgue" "lebesgue.\ s \ \" shows "(indicator s :: _ \ real) integrable_on UNIV" -proof (cases "lmeasure s") +proof (cases "lebesgue.\ s") case (preal m) from lmeasure_finite_has_integral[OF `s \ sets lebesgue` this] show ?thesis unfolding integrable_on_def by auto qed (insert assms, auto) @@ -314,7 +319,7 @@ qed lemma has_integral_lmeasure: assumes "((indicator s :: _\real) has_integral m) UNIV" - shows "lmeasure s = Real m" + shows "lebesgue.\ s = Real m" proof (intro lmeasure_iff_LIMSEQ[THEN iffD2]) let ?I = "indicator :: 'a set \ 'a \ real" show "s \ sets lebesgue" using has_integral_lebesgue[OF assms] . @@ -339,37 +344,37 @@ qed lemma has_integral_iff_lmeasure: - "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ lmeasure A = Real m)" + "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m)" proof assume "(indicator A has_integral m) UNIV" with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this] - show "A \ sets lebesgue \ 0 \ m \ lmeasure A = Real m" + show "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m" by (auto intro: has_integral_nonneg) next - assume "A \ sets lebesgue \ 0 \ m \ lmeasure A = Real m" + assume "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m" then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto qed lemma lmeasure_eq_integral: assumes "(indicator s::_\real) integrable_on UNIV" - shows "lmeasure s = Real (integral UNIV (indicator s))" + shows "lebesgue.\ s = Real (integral UNIV (indicator s))" using assms unfolding integrable_on_def proof safe fix y :: real assume "(indicator s has_integral y) UNIV" from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this] - show "lmeasure s = Real (integral UNIV (indicator s))" by simp + show "lebesgue.\ s = Real (integral UNIV (indicator s))" by simp qed lemma lebesgue_simple_function_indicator: fixes f::"'a::ordered_euclidean_space \ pextreal" - assumes f:"lebesgue.simple_function f" + assumes f:"simple_function lebesgue f" shows "f = (\x. (\y \ f ` UNIV. y * indicator (f -` {y}) x))" - apply(rule,subst lebesgue.simple_function_indicator_representation[OF f]) by auto + by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto lemma integral_eq_lmeasure: - "(indicator s::_\real) integrable_on UNIV \ integral UNIV (indicator s) = real (lmeasure s)" + "(indicator s::_\real) integrable_on UNIV \ integral UNIV (indicator s) = real (lebesgue.\ s)" by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg) -lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "lmeasure s \ \" +lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "lebesgue.\ s \ \" using lmeasure_eq_integral[OF assms] by auto lemma negligible_iff_lebesgue_null_sets: @@ -402,14 +407,13 @@ shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) -lemma lmeasure_UNIV[intro]: "lmeasure (UNIV::'a::ordered_euclidean_space set) = \" - unfolding lmeasure_def SUP_\ -proof (intro allI impI) +lemma lmeasure_UNIV[intro]: "lebesgue.\ (UNIV::'a::ordered_euclidean_space set) = \" +proof (simp add: lebesgue_def SUP_\, intro allI impI) fix x assume "x < \" then obtain r where r: "x = Real r" "0 \ r" by (cases x) auto then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto - show "\i\UNIV. x < Real (integral (cube i) (indicator UNIV::'a\real))" - proof (intro bexI[of _ n]) + show "\i. x < Real (integral (cube i) (indicator UNIV::'a\real))" + proof (intro exI[of _ n]) have [simp]: "indicator UNIV = (\x. 1)" by (auto simp: fun_eq_iff) { fix m :: nat assume "0 < m" then have "real n \ (\x Real (integral (cube n) (indicator UNIV::'a\real))" by (auto simp: real_eq_of_nat[symmetric] cube_def content_closed_interval_cases) finally show "x < Real (integral (cube n) (indicator UNIV::'a\real))" . - qed auto + qed qed lemma fixes a b ::"'a::ordered_euclidean_space" - shows lmeasure_atLeastAtMost[simp]: "lmeasure {a..b} = Real (content {a..b})" + shows lmeasure_atLeastAtMost[simp]: "lebesgue.\ {a..b} = Real (content {a..b})" proof - have "(indicator (UNIV \ {a..b})::_\real) integrable_on UNIV" unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw) @@ -453,7 +457,7 @@ qed lemma lmeasure_singleton[simp]: - fixes a :: "'a::ordered_euclidean_space" shows "lmeasure {a} = 0" + fixes a :: "'a::ordered_euclidean_space" shows "lebesgue.\ {a} = 0" using lmeasure_atLeastAtMost[of a a] by simp declare content_real[simp] @@ -461,74 +465,97 @@ lemma fixes a b :: real shows lmeasure_real_greaterThanAtMost[simp]: - "lmeasure {a <.. b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a <.. b} = Real (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lmeasure {a <.. b} = lmeasure {a .. b} - lmeasure {a}" + then have "lebesgue.\ {a <.. b} = lebesgue.\ {a .. b} - lebesgue.\ {a}" by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lmeasure]) + (auto intro!: arg_cong[where f=lebesgue.\]) then show ?thesis by auto qed auto lemma fixes a b :: real shows lmeasure_real_atLeastLessThan[simp]: - "lmeasure {a ..< b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a ..< b} = Real (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lmeasure {a ..< b} = lmeasure {a .. b} - lmeasure {b}" + then have "lebesgue.\ {a ..< b} = lebesgue.\ {a .. b} - lebesgue.\ {b}" by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lmeasure]) + (auto intro!: arg_cong[where f=lebesgue.\]) then show ?thesis by auto qed auto lemma fixes a b :: real shows lmeasure_real_greaterThanLessThan[simp]: - "lmeasure {a <..< b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a <..< b} = Real (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lmeasure {a <..< b} = lmeasure {a <.. b} - lmeasure {b}" + then have "lebesgue.\ {a <..< b} = lebesgue.\ {a <.. b} - lebesgue.\ {b}" by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lmeasure]) + (auto intro!: arg_cong[where f=lebesgue.\]) then show ?thesis by auto qed auto -interpretation borel: measure_space borel lmeasure -proof - show "countably_additive borel lmeasure" - using lebesgue.ca unfolding countably_additive_def - apply safe apply (erule_tac x=A in allE) by auto -qed auto +definition "lborel = lebesgue \ sets := sets borel \" + +lemma + shows space_lborel[simp]: "space lborel = UNIV" + and sets_lborel[simp]: "sets lborel = sets borel" + and measure_lborel[simp]: "measure lborel = lebesgue.\" + and measurable_lborel[simp]: "measurable lborel = measurable borel" + by (simp_all add: measurable_def_raw lborel_def) -interpretation borel: sigma_finite_measure borel lmeasure -proof (default, intro conjI exI[of _ "\n. cube n"]) - show "range cube \ sets borel" by (auto intro: borel_closed) - { fix x have "\n. x\cube n" using mem_big_cube by auto } - thus "(\i. cube i) = space borel" by auto - show "\i. lmeasure (cube i) \ \" unfolding cube_def by auto -qed +interpretation lborel: measure_space lborel + where "space lborel = UNIV" + and "sets lborel = sets borel" + and "measure lborel = lebesgue.\" + and "measurable lborel = measurable borel" +proof - + show "measure_space lborel" + proof + show "countably_additive lborel (measure lborel)" + using lebesgue.ca unfolding countably_additive_def lborel_def + apply safe apply (erule_tac x=A in allE) by auto + qed (auto simp: lborel_def) +qed simp_all -interpretation lebesgue: sigma_finite_measure lebesgue lmeasure +interpretation lborel: sigma_finite_measure lborel + where "space lborel = UNIV" + and "sets lborel = sets borel" + and "measure lborel = lebesgue.\" + and "measurable lborel = measurable borel" +proof - + show "sigma_finite_measure lborel" + proof (default, intro conjI exI[of _ "\n. cube n"]) + show "range cube \ sets lborel" by (auto intro: borel_closed) + { fix x have "\n. x\cube n" using mem_big_cube by auto } + thus "(\i. cube i) = space lborel" by auto + show "\i. measure lborel (cube i) \ \" by (simp add: cube_def) + qed +qed simp_all + +interpretation lebesgue: sigma_finite_measure lebesgue proof - from borel.sigma_finite guess A .. + from lborel.sigma_finite guess A .. moreover then have "range A \ sets lebesgue" using lebesgueI_borel by blast - ultimately show "\A::nat \ 'b set. range A \ sets lebesgue \ (\i. A i) = space lebesgue \ (\i. lmeasure (A i) \ \)" + ultimately show "\A::nat \ 'b set. range A \ sets lebesgue \ (\i. A i) = space lebesgue \ (\i. lebesgue.\ (A i) \ \)" by auto qed lemma simple_function_has_integral: fixes f::"'a::ordered_euclidean_space \ pextreal" - assumes f:"lebesgue.simple_function f" + assumes f:"simple_function lebesgue f" and f':"\x. f x \ \" - and om:"\x\range f. lmeasure (f -` {x} \ UNIV) = \ \ x = 0" - shows "((\x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV" - unfolding lebesgue.simple_integral_def + and om:"\x\range f. lebesgue.\ (f -` {x} \ UNIV) = \ \ x = 0" + shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" + unfolding simple_integral_def apply(subst lebesgue_simple_function_indicator[OF f]) proof - case goal1 have *:"\x. \y\range f. y * indicator (f -` {y}) x \ \" - "\x\range f. x * lmeasure (f -` {x} \ UNIV) \ \" + "\x\range f. x * lebesgue.\ (f -` {x} \ UNIV) \ \" using f' om unfolding indicator_def by auto show ?case unfolding space_lebesgue real_of_pextreal_setsum'[OF *(1),THEN sym] unfolding real_of_pextreal_setsum'[OF *(2),THEN sym] @@ -536,11 +563,11 @@ apply(rule has_integral_setsum) proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD) fix y::'a show "((\x. real (f y * indicator (f -` {f y}) x)) has_integral - real (f y * lmeasure (f -` {f y} \ UNIV))) UNIV" + real (f y * lebesgue.\ (f -` {f y} \ UNIV))) UNIV" proof(cases "f y = 0") case False have mea:"(indicator (f -` {f y}) ::_\real) integrable_on UNIV" apply(rule lmeasure_finite_integrable) - using assms unfolding lebesgue.simple_function_def using False by auto + using assms unfolding simple_function_def using False by auto have *:"\x. real (indicator (f -` {f y}) x::pextreal) = (indicator (f -` {f y}) x)" by (auto simp: indicator_def) show ?thesis unfolding real_of_pextreal_mult[THEN sym] @@ -558,31 +585,31 @@ lemma simple_function_has_integral': fixes f::"'a::ordered_euclidean_space \ pextreal" - assumes f:"lebesgue.simple_function f" - and i: "lebesgue.simple_integral f \ \" - shows "((\x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV" + assumes f:"simple_function lebesgue f" + and i: "integral\<^isup>S lebesgue f \ \" + shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" proof- let ?f = "\x. if f x = \ then 0 else f x" { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this have **:"{x. f x \ ?f x} = f -` {\}" by auto - have **:"lmeasure {x\space lebesgue. f x \ ?f x} = 0" + have **:"lebesgue.\ {x\space lebesgue. f x \ ?f x} = 0" using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**) show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **]) apply(rule lebesgue.simple_function_compose1[OF f]) unfolding * defer apply(rule simple_function_has_integral) proof- - show "lebesgue.simple_function ?f" + show "simple_function lebesgue ?f" using lebesgue.simple_function_compose1[OF f] . show "\x. ?f x \ \" by auto - show "\x\range ?f. lmeasure (?f -` {x} \ UNIV) = \ \ x = 0" + show "\x\range ?f. lebesgue.\ (?f -` {x} \ UNIV) = \ \ x = 0" proof (safe, simp, safe, rule ccontr) fix y assume "f y \ \" "f y \ 0" hence "(\x. if f x = \ then 0 else f x) -` {if f y = \ then 0 else f y} = f -` {f y}" by (auto split: split_if_asm) - moreover assume "lmeasure ((\x. if f x = \ then 0 else f x) -` {if f y = \ then 0 else f y}) = \" - ultimately have "lmeasure (f -` {f y}) = \" by simp + moreover assume "lebesgue.\ ((\x. if f x = \ then 0 else f x) -` {if f y = \ then 0 else f y}) = \" + ultimately have "lebesgue.\ (f -` {f y}) = \" by simp moreover - have "f y * lmeasure (f -` {f y}) \ \" using i f - unfolding lebesgue.simple_integral_def setsum_\ lebesgue.simple_function_def + have "f y * lebesgue.\ (f -` {f y}) \ \" using i f + unfolding simple_integral_def setsum_\ simple_function_def by auto ultimately have "f y = 0" by (auto split: split_if_asm) then show False using `f y \ 0` by simp @@ -595,7 +622,7 @@ assumes i: "\i. f i \ borel_measurable M" and mono: "\x. mono (\n. f n x)" and lim: "\x. (\i. f i x) ----> u x" shows "u \ borel_measurable M" - and "(\i. positive_integral (f i)) ----> positive_integral u" (is ?ilim) + and "(\i. integral\<^isup>P M (f i)) ----> integral\<^isup>P M u" (is ?ilim) proof - from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u] show ?ilim using mono lim i by auto @@ -609,19 +636,19 @@ lemma positive_integral_has_integral: fixes f::"'a::ordered_euclidean_space => pextreal" assumes f:"f \ borel_measurable lebesgue" - and int_om:"lebesgue.positive_integral f \ \" + and int_om:"integral\<^isup>P lebesgue f \ \" and f_om:"\x. f x \ \" (* TODO: remove this *) - shows "((\x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV" -proof- let ?i = "lebesgue.positive_integral f" + shows "((\x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV" +proof- let ?i = "integral\<^isup>P lebesgue f" from lebesgue.borel_measurable_implies_simple_function_sequence[OF f] guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2) let ?u = "\i x. real (u i x)" and ?f = "\x. real (f x)" - have u_simple:"\k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)" + have u_simple:"\k. integral\<^isup>S lebesgue (u k) = integral\<^isup>P lebesgue (u k)" apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) .. - have int_u_le:"\k. lebesgue.simple_integral (u k) \ lebesgue.positive_integral f" + have int_u_le:"\k. integral\<^isup>S lebesgue (u k) \ integral\<^isup>P lebesgue f" unfolding u_simple apply(rule lebesgue.positive_integral_mono) using isoton_Sup[OF u(3)] unfolding le_fun_def by auto - have u_int_om:"\i. lebesgue.simple_integral (u i) \ \" + have u_int_om:"\i. integral\<^isup>S lebesgue (u i) \ \" proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed note u_int = simple_function_has_integral'[OF u(1) this] @@ -633,17 +660,17 @@ prefer 3 apply(subst Real_real') defer apply(subst Real_real') using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto next case goal3 - show ?case apply(rule bounded_realI[where B="real (lebesgue.positive_integral f)"]) + show ?case apply(rule bounded_realI[where B="real (integral\<^isup>P lebesgue f)"]) apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int) unfolding integral_unique[OF u_int] defer apply(rule real_of_pextreal_mono[OF _ int_u_le]) using u int_om by auto qed note int = conjunctD2[OF this] - have "(\i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple + have "(\i. integral\<^isup>S lebesgue (u i)) ----> ?i" unfolding u_simple apply(rule lebesgue.positive_integral_monotone_convergence(2)) apply(rule lebesgue.borel_measurable_simple_function[OF u(1)]) using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto - hence "(\i. real (lebesgue.simple_integral (u i))) ----> real ?i" apply- + hence "(\i. real (integral\<^isup>S lebesgue (u i))) ----> real ?i" apply- apply(subst lim_Real[THEN sym]) prefer 3 apply(subst Real_real') defer apply(subst Real_real') using u f_om int_om u_int_om by auto @@ -653,12 +680,12 @@ lemma lebesgue_integral_has_integral: fixes f::"'a::ordered_euclidean_space => real" - assumes f:"lebesgue.integrable f" - shows "(f has_integral (lebesgue.integral f)) UNIV" + assumes f:"integrable lebesgue f" + shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV" proof- let ?n = "\x. - min (f x) 0" and ?p = "\x. max (f x) 0" have *:"f = (\x. ?p x - ?n x)" apply rule by auto - note f = lebesgue.integrableD[OF f] - show ?thesis unfolding lebesgue.integral_def apply(subst *) + note f = integrableD[OF f] + show ?thesis unfolding lebesgue_integral_def apply(subst *) proof(rule has_integral_sub) case goal1 have *:"\x. Real (f x) \ \" by auto note lebesgue.borel_measurable_Real[OF f(1)] @@ -674,27 +701,27 @@ qed lemma lebesgue_positive_integral_eq_borel: - "f \ borel_measurable borel \ lebesgue.positive_integral f = borel.positive_integral f " + "f \ borel_measurable borel \ integral\<^isup>P lebesgue f = integral\<^isup>P lborel f" by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default lemma lebesgue_integral_eq_borel: assumes "f \ borel_measurable borel" - shows "lebesgue.integrable f = borel.integrable f" (is ?P) - and "lebesgue.integral f = borel.integral f" (is ?I) + shows "integrable lebesgue f \ integrable lborel f" (is ?P) + and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I) proof - - have *: "sigma_algebra borel" by default - have "sets borel \ sets lebesgue" by auto - from lebesgue.integral_subalgebra[OF assms this _ *] + have *: "sigma_algebra lborel" by default + have "sets lborel \ sets lebesgue" by auto + from lebesgue.integral_subalgebra[of f lborel, OF _ this _ _ *] assms show ?P ?I by auto qed lemma borel_integral_has_integral: fixes f::"'a::ordered_euclidean_space => real" - assumes f:"borel.integrable f" - shows "(f has_integral (borel.integral f)) UNIV" + assumes f:"integrable lborel f" + shows "(f has_integral (integral\<^isup>L lborel f)) UNIV" proof - have borel: "f \ borel_measurable borel" - using f unfolding borel.integrable_def by auto + using f unfolding integrable_def by auto from f show ?thesis using lebesgue_integral_has_integral[of f] unfolding lebesgue_integral_eq_borel[OF borel] by simp @@ -708,11 +735,11 @@ using continuous_open_preimage[OF assms] unfolding vimage_def by auto lemma (in measure_space) integral_monotone_convergence_pos': - assumes i: "\i. integrable (f i)" and mono: "\x. mono (\n. f n x)" + assumes i: "\i. integrable M (f i)" and mono: "\x. mono (\n. f n x)" and pos: "\x i. 0 \ f i x" and lim: "\x. (\i. f i x) ----> u x" - and ilim: "(\i. integral (f i)) ----> x" - shows "integrable u \ integral u = x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + shows "integrable M u \ integral\<^isup>L M u = x" using integral_monotone_convergence_pos[OF assms] by auto definition e2p :: "'a::ordered_euclidean_space \ (nat \ real)" where @@ -751,53 +778,68 @@ thus "x \ e2p ` A" unfolding image_iff apply(rule_tac x="p2e x" in bexI) apply(subst e2p_p2e) by auto qed -interpretation borel_product: product_sigma_finite "\x. borel::real algebra" "\x. lmeasure" +interpretation lborel_product: product_sigma_finite "\x. lborel::real measure_space" by default -lemma cube_subset_Suc[intro]: "cube n \ cube (Suc n)" - unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto - -lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" - unfolding Pi_def by auto +interpretation lborel_space: finite_product_sigma_finite "\x. lborel::real measure_space" "{.." + and "measurable lborel = measurable borel" +proof - + show "finite_product_sigma_finite (\x. lborel::real measure_space) {.. measurable \ space = UNIV::'a set, sets = range lessThan \ - (product_algebra - (\x. \ space = UNIV::real set, sets = range lessThan \) - {.. measurable ?E ?P") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "e2p \ space ?E \ space ?P" by (auto simp: e2p_def) - fix A assume "A \ sets ?P" - then obtain E where A: "A = (\\<^isub>E i\{.. {.. (range lessThan)" - by (auto elim!: product_algebraE) - then have "\i\{..xs. E i = {..< xs}" by auto - from this[THEN bchoice] guess xs .. - then have A_eq: "A = (\\<^isub>E i\{..\ i. xs i) :: 'a}" - using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq - euclidean_eq[where 'a='a] eucl_less[where 'a='a]) - then show "e2p -` A \ space ?E \ sets ?E" by simp +lemma sets_product_borel: + assumes [intro]: "finite I" + shows "sets (\\<^isub>M i\I. + \ space = UNIV::real set, sets = range lessThan, measure = lebesgue.\ \) = + sets (\\<^isub>M i\I. lborel)" (is "sets ?G = _") +proof - + have "sets ?G = sets (\\<^isub>M i\I. + sigma \ space = UNIV::real set, sets = range lessThan, measure = lebesgue.\ \)" + by (subst sigma_product_algebra_sigma_eq[of I "\_ i. {.. measurable (borel::'a algebra) - (sigma (product_algebra (\x. borel :: real algebra) {..\<^isub>M i\{.. measurable ?E ?P") +proof - + let ?B = "\ space = UNIV::real set, sets = range lessThan, measure = lebesgue.\ \" + let ?G = "product_algebra_generator {.._. ?B)" + have "e2p \ measurable ?E (sigma ?G)" + proof (rule borel.measurable_sigma) + show "e2p \ space ?E \ space ?G" by (auto simp: e2p_def) + fix A assume "A \ sets ?G" + then obtain E where A: "A = (\\<^isub>E i\{.. {.. (range lessThan)" + by (auto elim!: product_algebraE simp: ) + then have "\i\{..xs. E i = {..< xs}" by auto + from this[THEN bchoice] guess xs .. + then have A_eq: "A = (\\<^isub>E i\{..\ i. xs i) :: 'a}" + using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq + euclidean_eq[where 'a='a] eucl_less[where 'a='a]) + then show "e2p -` A \ space ?E \ sets ?E" by simp + qed (auto simp: product_algebra_generator_def) + with sets_product_borel[of "{.. measurable - (product_algebra - (\x. \ space = UNIV::real set, sets = range lessThan \) - {.. space = UNIV::'a set, sets = range lessThan \" - (is "p2e \ measurable ?P ?E") -proof (unfold measurable_def, intro CollectI conjI ballI) +lemma measurable_p2e: + "p2e \ measurable (\\<^isub>M i\{.. measurable ?P _") + unfolding borel_eq_lessThan +proof (intro lborel_space.measurable_sigma) + let ?E = "\ space = UNIV :: 'a set, sets = range lessThan \" show "p2e \ space ?P \ space ?E" by simp fix A assume "A \ sets ?E" then obtain x where "A = {.. sets borel" - shows "lmeasure A = borel_product.product_measure {.. real) set)" + shows "lebesgue.\ A = lborel_space.\ TYPE('a) (e2p ` A)" (is "_ = ?m A") proof (rule measure_unique_Int_stable[where X=A and A=cube]) - interpret fprod: finite_product_sigma_finite "\x. borel :: real algebra" "\x. lmeasure" "{.. space = UNIV :: 'a set, sets = range (\(a,b). {a..b}) \" (is "Int_stable ?E" ) using Int_stable_cuboids' . - show "borel = sigma ?E" using borel_eq_atLeastAtMost . - show "\i. lmeasure (cube i) \ \" unfolding cube_def by auto - show "\X. X \ sets ?E \ - lmeasure X = borel_product.product_measure {.. real) set)" + have [simp]: "sigma ?E = borel" using borel_eq_atLeastAtMost .. + show "\i. lebesgue.\ (cube i) \ \" unfolding cube_def by auto + show "\X. X \ sets ?E \ lebesgue.\ X = ?m X" proof- case goal1 then obtain a b where X:"X = {a..b}" by auto { presume *:"X \ {} \ ?case" show ?case apply(cases,rule *,assumption) by auto } @@ -861,12 +893,12 @@ show "x \ Pi\<^isub>E {..x X = (\xii (XX i))" apply(rule setprod_cong2) unfolding XX_def lmeasure_atLeastAtMost apply(subst content_real) using X' by auto - also have "... = borel_product.product_measure {..x. \xa. x \ cube xa" apply(rule_tac x=x in mem_big_cube) by fastsimp thus "cube \ space \space = UNIV, sets = range (\(a, b). {a..b})\" apply-apply(rule isotoneI) apply(rule cube_subset_Suc) by auto - show "A \ sets borel " by fact - show "measure_space borel lmeasure" by default - show "measure_space borel - (\a::'a set. finite_product_sigma_finite.measure (\x. borel) (\x. lmeasure) {.. real) \ 'a) \ measurable fprod.P borel" by (rule measurable_p2e) - fix A :: "'a set" assume "A \ sets borel" - show "fprod.measure (e2p ` A) = fprod.measure (p2e -` A \ space fprod.P)" + show "A \ sets (sigma ?E)" using assms by simp + have "measure_space lborel" by default + then show "measure_space \ space = space ?E, sets = sets (sigma ?E), measure = measure lebesgue\" + unfolding lebesgue_def lborel_def by simp + let ?M = "\ space = space ?E, sets = sets (sigma ?E), measure = ?m \" + show "measure_space ?M" + proof (rule lborel_space.measure_space_vimage) + show "sigma_algebra ?M" by (rule lborel.sigma_algebra_cong) auto + show "p2e \ measurable (\\<^isub>M i\{.. sets ?M" + show "measure ?M A = lborel_space.\ TYPE('a) (p2e -` A \ space (\\<^isub>M i\{.. _) = extensional {.. pextreal" assumes f: "f \ borel_measurable borel" - shows "borel.positive_integral f = - borel_product.product_positive_integral {.. p2e)" -proof- def U \ "extensional {.. real) set" - interpret fprod: finite_product_sigma_finite "\x. borel" "\x. lmeasure" "{.. _" "(\x. f (p2e x))", unfolded p2e_e2p]) - show "(e2p :: 'a \ _) \ measurable borel fprod.P" by (rule measurable_e2p) - show "sigma_algebra fprod.P" by default - from measurable_comp[OF measurable_p2e f] - show "(\x. f (p2e x)) \ borel_measurable fprod.P" by (simp add: comp_def) - let "?L A" = "lmeasure ((e2p::'a \ (nat \ real)) -` A \ space borel)" - show "measure_space.positive_integral fprod.P ?L (\x. f (p2e x)) = - fprod.positive_integral (f \ p2e)" - unfolding comp_def - proof (rule fprod.positive_integral_cong_measure) - fix A :: "(nat \ real) set" assume "A \ sets fprod.P" - then have A: "(e2p::'a \ (nat \ real)) -` A \ space borel \ sets borel" - by (rule measurable_sets[OF measurable_e2p]) - have [simp]: "A \ extensional {.. sets fprod.P`[THEN fprod.sets_into_space] by auto - show "?L A = fprod.measure A" - unfolding lmeasure_measure_eq_borel_prod[OF A] - by (simp add: range_e2p) - qed - qed + shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(lborel_space.P TYPE('a))" +proof (rule lborel.positive_integral_vimage[symmetric, of _ "e2p :: 'a \ _" "(\x. f (p2e x))", unfolded p2e_e2p]) + show "(e2p :: 'a \ _) \ measurable borel (lborel_space.P TYPE('a))" by (rule measurable_e2p) + show "sigma_algebra (lborel_space.P TYPE('a))" by default + from measurable_comp[OF measurable_p2e f] + show "(\x. f (p2e x)) \ borel_measurable (lborel_space.P TYPE('a))" by (simp add: comp_def) + let "?L A" = "lebesgue.\ ((e2p::'a \ (nat \ real)) -` A \ UNIV)" + fix A :: "(nat \ real) set" assume "A \ sets (lborel_space.P TYPE('a))" + then have A: "(e2p::'a \ (nat \ real)) -` A \ space borel \ sets borel" + by (rule measurable_sets[OF measurable_e2p]) + have [simp]: "A \ extensional {.. sets (lborel_space.P TYPE('a))`[THEN lborel_space.sets_into_space] by auto + show "lborel_space.\ TYPE('a) A = ?L A" + using lmeasure_measure_eq_borel_prod[OF A] by (simp add: range_e2p) qed lemma borel_fubini: fixes f :: "'a::ordered_euclidean_space \ real" assumes f: "f \ borel_measurable borel" - shows "borel.integral f = borel_product.product_integral {.. p2e)" -proof- interpret fprod: finite_product_sigma_finite "\x. borel" "\x. lmeasure" "{..L lborel f = \x. f (p2e x) \(lborel_space.P TYPE('a))" +proof - have 1:"(\x. Real (f x)) \ borel_measurable borel" using f by auto have 2:"(\x. Real (- f x)) \ borel_measurable borel" using f by auto - show ?thesis unfolding fprod.integral_def borel.integral_def + show ?thesis unfolding lebesgue_integral_def unfolding borel_fubini_positiv_integral[OF 1] borel_fubini_positiv_integral[OF 2] unfolding o_def .. qed diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Measure.thy --- a/src/HOL/Probability/Measure.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Measure.thy Wed Feb 02 12:34:45 2011 +0100 @@ -4,6 +4,18 @@ imports Caratheodory begin +lemma measure_algebra_more[simp]: + "\ space = A, sets = B, \ = algebra.more M \ \ measure := m \ = + \ space = A, sets = B, \ = algebra.more (M \ measure := m \) \" + by (cases M) simp + +lemma measure_algebra_more_eq[simp]: + "\X. measure \ space = T, sets = A, \ = algebra.more X \ = measure X" + unfolding measure_space.splits by simp + +lemma measure_sigma[simp]: "measure (sigma A) = measure A" + unfolding sigma_def by simp + lemma inj_on_image_eq_iff: assumes "inj_on f S" assumes "A \ S" "B \ S" @@ -53,24 +65,34 @@ with ca assms show ?thesis by (simp add: countably_additive_def) qed +lemma (in sigma_algebra) sigma_algebra_cong: + assumes "space N = space M" "sets N = sets M" + shows "sigma_algebra N" + by default (insert sets_into_space, auto simp: assms) + lemma (in measure_space) measure_space_cong: - assumes "\A. A \ sets M \ \ A = \ A" - shows "measure_space M \" -proof - show "\ {} = 0" using assms by auto - show "countably_additive M \" unfolding countably_additive_def - proof safe - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" - then have "\i. A i \ sets M" "(UNION UNIV A) \ sets M" by auto - from this[THEN assms] measure_countably_additive[OF A] - show "(\\<^isub>\n. \ (A n)) = \ (UNION UNIV A)" by simp + assumes "\A. A \ sets M \ measure N A = \ A" "space N = space M" "sets N = sets M" + shows "measure_space N" +proof - + interpret N: sigma_algebra N by (intro sigma_algebra_cong assms) + show ?thesis + proof + show "measure N {} = 0" using assms by auto + show "countably_additive N (measure N)" unfolding countably_additive_def + proof safe + fix A :: "nat \ 'a set" assume A: "range A \ sets N" "disjoint_family A" + then have "\i. A i \ sets M" "(UNION UNIV A) \ sets M" unfolding assms by auto + from measure_countably_additive[of A] A this[THEN assms(1)] + show "(\\<^isub>\n. measure N (A n)) = measure N (UNION UNIV A)" + unfolding assms by simp + qed qed qed lemma (in measure_space) additive: "additive M \" proof (rule algebra.countably_additive_additive [OF _ _ ca]) show "algebra M" by default - show "positive \" by (simp add: positive_def) + show "positive M \" by (simp add: positive_def) qed lemma (in measure_space) measure_additive: @@ -358,12 +380,16 @@ finally show ?thesis by simp qed -lemma (in sigma_algebra) finite_additivity_sufficient: - assumes fin: "finite (space M)" and pos: "positive \" and add: "additive M \" - shows "measure_space M \" -proof - show [simp]: "\ {} = 0" using pos by (simp add: positive_def) - show "countably_additive M \" +lemma finite_additivity_sufficient: + assumes "sigma_algebra M" + assumes fin: "finite (space M)" and pos: "positive M (measure M)" and add: "additive M (measure M)" + shows "measure_space M" +proof - + interpret sigma_algebra M by fact + show ?thesis + proof + show [simp]: "measure M {} = 0" using pos by (simp add: positive_def) + show "countably_additive M (measure M)" proof (auto simp add: countably_additive_def) fix A :: "nat \ 'a set" assume A: "range A \ sets M" @@ -391,15 +417,15 @@ by blast qed then obtain N where N: "\m\N. A m = {}" by blast - then have "\m\N. \ (A m) = 0" by simp - then have "(\\<^isub>\ n. \ (A n)) = setsum (\m. \ (A m)) {..m\N. measure M (A m) = 0" by simp + then have "(\\<^isub>\ n. measure M (A n)) = setsum (\m. measure M (A m)) {.. (\ii (A n \ (\ x (A n) + \ (\ i (\ x i (\ x (\i. A i)" + also have "... = measure M (\i. A i)" proof - have "(\ ii. A i)" using N by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE) thus ?thesis by simp qed - finally show "(\\<^isub>\ n. \ (A n)) = \ (\i. A i)" . + finally show "(\\<^isub>\ n. measure M (A n)) = measure M (\i. A i)" . qed + qed qed lemma (in measure_space) measure_setsum_split: @@ -525,95 +552,75 @@ qed qed -lemma True -proof - fix x a b :: nat - have "\x a b :: int. x dvd a \ x dvd (a + b) \ x dvd b" - by (metis dvd_mult_div_cancel zadd_commute zdvd_reduce) - then have "x dvd a \ x dvd (a + b) \ x dvd b" - unfolding zdvd_int[of x] zadd_int[symmetric] . -qed - lemma measure_unique_Int_stable: - fixes M E :: "'a algebra" and A :: "nat \ 'a set" - assumes "Int_stable E" "M = sigma E" - and A: "range A \ sets E" "A \ space E" - and ms: "measure_space M \" "measure_space M \" + fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \ 'a set" + assumes "Int_stable E" + and A: "range A \ sets E" "A \ space E" + and M: "measure_space \space = space E, sets = sets (sigma E), measure = \\" (is "measure_space ?M") + and N: "measure_space \space = space E, sets = sets (sigma E), measure = \\" (is "measure_space ?N") and eq: "\X. X \ sets E \ \ X = \ X" and finite: "\i. \ (A i) \ \" - assumes "X \ sets M" + assumes "X \ sets (sigma E)" shows "\ X = \ X" proof - - let "?D F" = "{D. D \ sets M \ \ (F \ D) = \ (F \ D)}" - interpret M: measure_space M \ by fact - interpret M': measure_space M \ by fact - have "space E = space M" - using `M = sigma E` by simp - have sets_E: "sets E \ Pow (space E)" - proof - fix X assume "X \ sets E" - then have "X \ sets M" unfolding `M = sigma E` - unfolding sigma_def by (auto intro!: sigma_sets.Basic) - with M.sets_into_space show "X \ Pow (space E)" - unfolding `space E = space M` by auto - qed - have A': "range A \ sets M" using `M = sigma E` A - by (auto simp: sets_sigma intro!: sigma_sets.Basic) + let "?D F" = "{D. D \ sets (sigma E) \ \ (F \ D) = \ (F \ D)}" + interpret M: measure_space ?M + where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \" by (simp_all add: M) + interpret N: measure_space ?N + where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \" by (simp_all add: N) { fix F assume "F \ sets E" and "\ F \ \" - then have [intro]: "F \ sets M" unfolding `M = sigma E` sets_sigma - by (intro sigma_sets.Basic) + then have [intro]: "F \ sets (sigma E)" by auto have "\ F \ \" using `\ F \ \` `F \ sets E` eq by simp interpret D: dynkin_system "\space=space E, sets=?D F\" proof (rule dynkin_systemI, simp_all) - fix A assume "A \ sets M \ \ (F \ A) = \ (F \ A)" - then show "A \ space E" - unfolding `space E = space M` using M.sets_into_space by auto + fix A assume "A \ sets (sigma E) \ \ (F \ A) = \ (F \ A)" + then show "A \ space E" using M.sets_into_space by auto next - have "F \ space E = F" using `F \ sets E` sets_E by auto - then show "space E \ sets M \ \ (F \ space E) = \ (F \ space E)" - unfolding `space E = space M` using `F \ sets E` eq by auto + have "F \ space E = F" using `F \ sets E` by auto + then show "\ (F \ space E) = \ (F \ space E)" + using `F \ sets E` eq by auto next - fix A assume *: "A \ sets M \ \ (F \ A) = \ (F \ A)" - then have **: "F \ (space M - A) = F - (F \ A)" - and [intro]: "F \ A \ sets M" - using `F \ sets E` sets_E `space E = space M` by auto - have "\ (F \ A) \ \ F" by (auto intro!: M'.measure_mono) + fix A assume *: "A \ sets (sigma E) \ \ (F \ A) = \ (F \ A)" + then have **: "F \ (space (sigma E) - A) = F - (F \ A)" + and [intro]: "F \ A \ sets (sigma E)" + using `F \ sets E` M.sets_into_space by auto + have "\ (F \ A) \ \ F" by (auto intro!: N.measure_mono) then have "\ (F \ A) \ \" using `\ F \ \` by auto have "\ (F \ A) \ \ F" by (auto intro!: M.measure_mono) then have "\ (F \ A) \ \" using `\ F \ \` by auto - then have "\ (F \ (space M - A)) = \ F - \ (F \ A)" unfolding ** - using `F \ A \ sets M` by (auto intro!: M.measure_Diff) + then have "\ (F \ (space (sigma E) - A)) = \ F - \ (F \ A)" unfolding ** + using `F \ A \ sets (sigma E)` by (auto intro!: M.measure_Diff) also have "\ = \ F - \ (F \ A)" using eq `F \ sets E` * by simp - also have "\ = \ (F \ (space M - A))" unfolding ** - using `F \ A \ sets M` `\ (F \ A) \ \` by (auto intro!: M'.measure_Diff[symmetric]) - finally show "space E - A \ sets M \ \ (F \ (space E - A)) = \ (F \ (space E - A))" - using `space E = space M` * by auto + also have "\ = \ (F \ (space (sigma E) - A))" unfolding ** + using `F \ A \ sets (sigma E)` `\ (F \ A) \ \` by (auto intro!: N.measure_Diff[symmetric]) + finally show "space E - A \ sets (sigma E) \ \ (F \ (space E - A)) = \ (F \ (space E - A))" + using * by auto next fix A :: "nat \ 'a set" - assume "disjoint_family A" "range A \ {X \ sets M. \ (F \ X) = \ (F \ X)}" - then have A: "range (\i. F \ A i) \ sets M" "F \ (\x. A x) = (\x. F \ A x)" - "disjoint_family (\i. F \ A i)" "\i. \ (F \ A i) = \ (F \ A i)" "range A \ sets M" - by ((fastsimp simp: disjoint_family_on_def)+) - then show "(\x. A x) \ sets M \ \ (F \ (\x. A x)) = \ (F \ (\x. A x))" + assume "disjoint_family A" "range A \ {X \ sets (sigma E). \ (F \ X) = \ (F \ X)}" + then have A: "range (\i. F \ A i) \ sets (sigma E)" "F \ (\x. A x) = (\x. F \ A x)" + "disjoint_family (\i. F \ A i)" "\i. \ (F \ A i) = \ (F \ A i)" "range A \ sets (sigma E)" + by (auto simp: disjoint_family_on_def subset_eq) + then show "(\x. A x) \ sets (sigma E) \ \ (F \ (\x. A x)) = \ (F \ (\x. A x))" by (auto simp: M.measure_countably_additive[symmetric] - M'.measure_countably_additive[symmetric] + N.measure_countably_additive[symmetric] simp del: UN_simps) qed - have *: "sigma E = \space = space E, sets = ?D F\" - using `M = sigma E` `F \ sets E` `Int_stable E` + have *: "sets (sigma E) = sets \space = space E, sets = ?D F\" + using `F \ sets E` `Int_stable E` by (intro D.dynkin_lemma) (auto simp add: sets_sigma Int_stable_def eq intro: sigma_sets.Basic) - have "\D. D \ sets M \ \ (F \ D) = \ (F \ D)" - unfolding `M = sigma E` by (auto simp: *) } + have "\D. D \ sets (sigma E) \ \ (F \ D) = \ (F \ D)" + by (subst (asm) *) auto } note * = this { fix i have "\ (A i \ X) = \ (A i \ X)" - using *[of "A i" X] `X \ sets M` A finite by auto } + using *[of "A i" X] `X \ sets (sigma E)` A finite by auto } moreover have "(\i. A i \ X) \ X" - using `X \ sets M` M.sets_into_space A `space E = space M` + using `X \ sets (sigma E)` M.sets_into_space A by (auto simp: isoton_def) then have "(\i. \ (A i \ X)) \ \ X" "(\i. \ (A i \ X)) \ \ X" - using `X \ sets M` A' by (auto intro!: M.measure_up M'.measure_up M.Int) + using `X \ sets (sigma E)` A by (auto intro!: M.measure_up N.measure_up M.Int simp: subset_eq) ultimately show ?thesis by (simp add: isoton_def) qed @@ -830,37 +837,38 @@ lemma (in measure_space) restricted_measure_space: assumes "S \ sets M" - shows "measure_space (restricted_space S) \" - (is "measure_space ?r \") + shows "measure_space (restricted_space S)" + (is "measure_space ?r") unfolding measure_space_def measure_space_axioms_def proof safe show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] . - show "\ {} = 0" by simp - show "countably_additive ?r \" + show "measure ?r {} = 0" by simp + + show "countably_additive ?r (measure ?r)" unfolding countably_additive_def proof safe fix A :: "nat \ 'a set" assume *: "range A \ sets ?r" and **: "disjoint_family A" from restriction_in_sets[OF assms *[simplified]] ** - show "(\\<^isub>\ n. \ (A n)) = \ (\i. A i)" + show "(\\<^isub>\ n. measure ?r (A n)) = measure ?r (\i. A i)" using measure_countably_additive by simp qed qed lemma (in measure_space) measure_space_vimage: - fixes M' :: "'b algebra" + fixes M' :: "('c, 'd) measure_space_scheme" assumes T: "sigma_algebra M'" "T \ measurable M M'" - and \: "\A. A \ sets M' \ \ A = \ (T -` A \ space M)" - shows "measure_space M' \" + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + shows "measure_space M'" proof - interpret M': sigma_algebra M' by fact show ?thesis proof - show "\ {} = 0" using \[of "{}"] by simp + show "measure M' {} = 0" using \[of "{}"] by simp - show "countably_additive M' \" - proof (intro countably_additive_def[THEN iffD2] allI impI) - fix A :: "nat \ 'b set" assume "range A \ sets M'" "disjoint_family A" + show "countably_additive M' (measure M')" + proof (intro countably_additiveI) + fix A :: "nat \ 'c set" assume "range A \ sets M'" "disjoint_family A" then have A: "\i. A i \ sets M'" "(\i. A i) \ sets M'" by auto then have *: "range (\i. T -` (A i) \ space M) \ sets M" using `T \ measurable M M'` by (auto simp: measurable_def) @@ -868,7 +876,7 @@ using * by blast moreover have **: "disjoint_family (\i. T -` A i \ space M)" using `disjoint_family A` by (auto simp: disjoint_family_on_def) - ultimately show "(\\<^isub>\ i. \ (A i)) = \ (\i. A i)" + ultimately show "(\\<^isub>\ i. measure M' (A i)) = measure M' (\i. A i)" using measure_countably_additive[OF _ **] A by (auto simp: comp_def vimage_UN \) qed @@ -877,14 +885,15 @@ lemma (in measure_space) measure_space_subalgebra: assumes "sigma_algebra N" and [simp]: "sets N \ sets M" "space N = space M" - shows "measure_space N \" + and measure[simp]: "\X. X \ sets N \ measure N X = measure M X" + shows "measure_space N" proof - interpret N: sigma_algebra N by fact show ?thesis proof from `sets N \ sets M` have "\A. range A \ sets N \ range A \ sets M" by blast - then show "countably_additive N \" - by (auto intro!: measure_countably_additive simp: countably_additive_def) + then show "countably_additive N (measure N)" + by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq) qed simp qed @@ -895,16 +904,16 @@ lemma (in sigma_finite_measure) restricted_sigma_finite_measure: assumes "S \ sets M" - shows "sigma_finite_measure (restricted_space S) \" - (is "sigma_finite_measure ?r _") + shows "sigma_finite_measure (restricted_space S)" + (is "sigma_finite_measure ?r") unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def proof safe - show "measure_space ?r \" using restricted_measure_space[OF assms] . + show "measure_space ?r" using restricted_measure_space[OF assms] . next obtain A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. \ (A i) \ \" using sigma_finite by auto - show "\A::nat \ 'a set. range A \ sets ?r \ (\i. A i) = space ?r \ (\i. \ (A i) \ \)" + show "\A::nat \ 'a set. range A \ sets ?r \ (\i. A i) = space ?r \ (\i. measure ?r (A i) \ \)" proof (safe intro!: exI[of _ "\i. A i \ S"] del: notI) fix i show "A i \ S \ sets ?r" @@ -919,22 +928,21 @@ have "\ (A i \ S) \ \ (A i)" using `range A \ sets M` `S \ sets M` by (auto intro!: measure_mono) also have "\ < \" using `\ (A i) \ \` by (auto simp: pextreal_less_\) - finally show "\ (A i \ S) \ \" by (auto simp: pextreal_less_\) + finally show "measure ?r (A i \ S) \ \" by (auto simp: pextreal_less_\) qed qed lemma (in sigma_finite_measure) sigma_finite_measure_cong: - assumes cong: "\A. A \ sets M \ \' A = \ A" - shows "sigma_finite_measure M \'" + assumes cong: "\A. A \ sets M \ measure M' A = \ A" "sets M' = sets M" "space M' = space M" + shows "sigma_finite_measure M'" proof - - interpret \': measure_space M \' - using cong by (rule measure_space_cong) + interpret M': measure_space M' by (intro measure_space_cong cong) from sigma_finite guess A .. note A = this then have "\i. A i \ sets M" by auto - with A have fin: "(\i. \' (A i) \ \)" using cong by auto + with A have fin: "(\i. measure M' (A i) \ \)" using cong by auto show ?thesis apply default - using A fin by auto + using A fin cong by auto qed lemma (in sigma_finite_measure) disjoint_sigma_finite: @@ -1110,20 +1118,20 @@ lemma (in finite_measure) restricted_finite_measure: assumes "S \ sets M" - shows "finite_measure (restricted_space S) \" - (is "finite_measure ?r _") + shows "finite_measure (restricted_space S)" + (is "finite_measure ?r") unfolding finite_measure_def finite_measure_axioms_def proof (safe del: notI) - show "measure_space ?r \" using restricted_measure_space[OF assms] . + show "measure_space ?r" using restricted_measure_space[OF assms] . next - show "\ (space ?r) \ \" using finite_measure[OF `S \ sets M`] by auto + show "measure ?r (space ?r) \ \" using finite_measure[OF `S \ sets M`] by auto qed lemma (in measure_space) restricted_to_finite_measure: assumes "S \ sets M" "\ S \ \" - shows "finite_measure (restricted_space S) \" + shows "finite_measure (restricted_space S)" proof - - have "measure_space (restricted_space S) \" + have "measure_space (restricted_space S)" using `S \ sets M` by (rule restricted_measure_space) then show ?thesis unfolding finite_measure_def finite_measure_axioms_def @@ -1218,105 +1226,104 @@ section {* Measure preserving *} -definition "measure_preserving A \ B \ = - {f \ measurable A B. (\y \ sets B. \ (f -` y \ space A) = \ y)}" +definition "measure_preserving A B = + {f \ measurable A B. (\y \ sets B. measure A (f -` y \ space A) = measure B y)}" lemma (in finite_measure) measure_preserving_lift: - fixes f :: "'a \ 'a2" and A :: "'a2 algebra" - assumes "algebra A" - assumes "finite_measure (sigma A) \" (is "finite_measure ?sA _") - assumes fmp: "f \ measure_preserving M \ A \" - shows "f \ measure_preserving M \ (sigma A) \" + fixes f :: "'a \ 'c" and A :: "('c, 'd) measure_space_scheme" + assumes "algebra A" "finite_measure (sigma A)" (is "finite_measure ?sA") + assumes fmp: "f \ measure_preserving M A" + shows "f \ measure_preserving M (sigma A)" proof - - interpret sA: finite_measure ?sA \ by fact + interpret sA: finite_measure ?sA by fact interpret A: algebra A by fact show ?thesis using fmp - proof (clarsimp simp add: measure_preserving_def) - assume fm: "f \ measurable M A" - and meq: "\y\sets A. \ (f -` y \ space M) = \ y" - have f12: "f \ measurable M ?sA" - using measurable_subset[OF A.sets_into_space] fm by auto - hence ffn: "f \ space M \ space A" - by (simp add: measurable_def) - have "space M \ f -` (space A)" - by auto (metis PiE ffn) - hence fveq [simp]: "(f -` (space A)) \ space M = space M" - by blast - { - fix y - assume y: "y \ sets ?sA" - have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def) - also have "\ \ {s . \ ((f -` s) \ space M) = \ s}" - proof (rule A.sigma_property_disjoint, auto) - fix x assume "x \ sets A" then show "\ (f -` x \ space M) = \ x" by (simp add: meq) - next - fix s - assume eq: "\ (f -` s \ space M) = \ s" and s: "s \ ?A" - then have s': "s \ sets ?sA" by (simp add: sigma_def) - show "\ (f -` (space A - s) \ space M) = \ (space A - s)" - using sA.finite_measure_compl[OF s'] - using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top] - by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq) - next - fix S - assume S0: "S 0 = {}" - and SSuc: "\n. S n \ S (Suc n)" - and rS1: "range S \ {s. \ (f -` s \ space M) = \ s}" - and "range S \ ?A" - hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) - have eq1: "\i. \ (f -` S i \ space M) = \ (S i)" - using rS1 by blast - have *: "(\n. \ (S n)) = (\n. \ (f -` S n \ space M))" - by (simp add: eq1) - have "(SUP n. ... n) = \ (\i. f -` S i \ space M)" - proof (rule measure_countable_increasing) - show "range (\i. f -` S i \ space M) \ sets M" - using f12 rS2 by (auto simp add: measurable_def) - show "f -` S 0 \ space M = {}" using S0 - by blast - show "\n. f -` S n \ space M \ f -` S (Suc n) \ space M" - using SSuc by auto - qed - also have "\ (\i. f -` S i \ space M) = \ (f -` (\i. S i) \ space M)" - by (simp add: vimage_UN) - finally have "(SUP n. \ (S n)) = \ (f -` (\i. S i) \ space M)" unfolding * . - moreover - have "(SUP n. \ (S n)) = \ (\i. S i)" - by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc]) - ultimately - show "\ (f -` (\i. S i) \ space M) = \ (\i. S i)" by simp - next - fix S :: "nat => 'a2 set" - assume dS: "disjoint_family S" - and rS1: "range S \ {s. \ (f -` s \ space M) = \ s}" - and "range S \ ?A" - hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) - have "\i. \ (f -` S i \ space M) = \ (S i)" - using rS1 by blast - hence *: "(\i. \ (S i)) = (\n. \ (f -` S n \ space M))" - by simp - have "psuminf ... = \ (\i. f -` S i \ space M)" - proof (rule measure_countably_additive) - show "range (\i. f -` S i \ space M) \ sets M" - using f12 rS2 by (auto simp add: measurable_def) - show "disjoint_family (\i. f -` S i \ space M)" using dS - by (auto simp add: disjoint_family_on_def) - qed - hence "(\\<^isub>\ i. \ (S i)) = \ (\i. f -` S i \ space M)" unfolding * . - with sA.measure_countably_additive [OF rS2 dS] - have "\ (\i. f -` S i \ space M) = \ (\i. S i)" - by simp - moreover have "\ (f -` (\i. S i) \ space M) = \ (\i. f -` S i \ space M)" - by (simp add: vimage_UN) - ultimately show "\ (f -` (\i. S i) \ space M) = \ (\i. S i)" - by metis + proof (clarsimp simp add: measure_preserving_def) + assume fm: "f \ measurable M A" + and "\y\sets A. \ (f -` y \ space M) = measure A y" + then have meq: "\y\sets A. \ (f -` y \ space M) = sA.\ y" + by simp + have f12: "f \ measurable M ?sA" + using measurable_subset[OF A.sets_into_space] fm by auto + hence ffn: "f \ space M \ space A" + by (simp add: measurable_def) + have "space M \ f -` (space A)" + by auto (metis PiE ffn) + hence fveq [simp]: "(f -` (space A)) \ space M = space M" + by blast + { + fix y + assume y: "y \ sets ?sA" + have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def) + also have "\ \ {s . \ ((f -` s) \ space M) = sA.\ s}" + proof (rule A.sigma_property_disjoint, safe) + fix x assume "x \ sets A" then show "\ (f -` x \ space M) = sA.\ x" by (simp add: meq) + next + fix s + assume eq: "\ (f -` s \ space M) = sA.\ s" and s: "s \ ?A" + then have s': "s \ sets ?sA" by (simp add: sigma_def) + show "\ (f -` (space A - s) \ space M) = measure ?sA (space A - s)" + using sA.finite_measure_compl[OF s'] + using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top] + by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq) + next + fix S + assume S0: "S 0 = {}" + and SSuc: "\n. S n \ S (Suc n)" + and rS1: "range S \ {s. \ (f -` s \ space M) = sA.\ s} \ ?A" + hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) + have eq1: "\i. \ (f -` S i \ space M) = sA.\ (S i)" + using rS1 by blast + have *: "(\n. sA.\ (S n)) = (\n. \ (f -` S n \ space M))" + by (simp add: eq1) + have "(SUP n. ... n) = \ (\i. f -` S i \ space M)" + proof (rule measure_countable_increasing) + show "range (\i. f -` S i \ space M) \ sets M" + using f12 rS2 by (auto simp add: measurable_def) + show "f -` S 0 \ space M = {}" using S0 + by blast + show "\n. f -` S n \ space M \ f -` S (Suc n) \ space M" + using SSuc by auto + qed + also have "\ (\i. f -` S i \ space M) = \ (f -` (\i. S i) \ space M)" + by (simp add: vimage_UN) + finally have "(SUP n. sA.\ (S n)) = \ (f -` (\i. S i) \ space M)" unfolding * . + moreover + have "(SUP n. sA.\ (S n)) = sA.\ (\i. S i)" + by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc]) + ultimately + show "\ (f -` (\i. S i) \ space M) = sA.\ (\i. S i)" by simp + next + fix S :: "nat \ 'c set" + assume dS: "disjoint_family S" + and rS1: "range S \ {s. \ (f -` s \ space M) = sA.\ s} \ ?A" + hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) + have "\i. \ (f -` S i \ space M) = sA.\ (S i)" + using rS1 by blast + hence *: "(\i. sA.\ (S i)) = (\n. \ (f -` S n \ space M))" + by simp + have "psuminf ... = \ (\i. f -` S i \ space M)" + proof (rule measure_countably_additive) + show "range (\i. f -` S i \ space M) \ sets M" + using f12 rS2 by (auto simp add: measurable_def) + show "disjoint_family (\i. f -` S i \ space M)" using dS + by (auto simp add: disjoint_family_on_def) qed - finally have "sets ?sA \ {s . \ ((f -` s) \ space M) = \ s}" . - hence "\ (f -` y \ space M) = \ y" using y by force - } - thus "f \ measurable M ?sA \ (\y\sets ?sA. \ (f -` y \ space M) = \ y)" - by (blast intro: f12) - qed + hence "(\\<^isub>\ i. sA.\ (S i)) = \ (\i. f -` S i \ space M)" unfolding * . + with sA.measure_countably_additive [OF rS2 dS] + have "\ (\i. f -` S i \ space M) = sA.\ (\i. S i)" + by simp + moreover have "\ (f -` (\i. S i) \ space M) = \ (\i. f -` S i \ space M)" + by (simp add: vimage_UN) + ultimately show "\ (f -` (\i. S i) \ space M) = sA.\ (\i. S i)" + by metis + qed + finally have "sets ?sA \ {s . \ ((f -` s) \ space M) = sA.\ s}" . + hence "\ (f -` y \ space M) = sA.\ y" using y by force + } + thus "f \ measurable M ?sA \ (\y\sets ?sA. \ (f -` y \ space M) = measure A y)" + by simp_all (blast intro: f12) + qed qed section "Finite spaces" @@ -1329,22 +1336,24 @@ by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) lemma finite_measure_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" and space: "\ (space M) \ \" - and add: "\A B. A\space M \ B\space M \ A \ B = {} \ \ (A \ B) = \ A + \ B" - and "\ {} = 0" - shows "finite_measure_space M \" + assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \ \" + and add: "\A B. A\space M \ B\space M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" + and "measure M {} = 0" + shows "finite_measure_space M" unfolding finite_measure_space_def finite_measure_space_axioms_def proof (intro allI impI conjI) - show "measure_space M \" - proof (rule sigma_algebra.finite_additivity_sufficient) - have *: "\space = space M, sets = sets M\ = M" by auto + show "measure_space M" + proof (rule finite_additivity_sufficient) + have *: "\space = space M, sets = Pow (space M), \ = algebra.more M\ = M" + unfolding assms(2)[symmetric] by (auto intro!: algebra.equality) show "sigma_algebra M" - using sigma_algebra_Pow[of "space M" "more M"] assms(2)[symmetric] by (simp add: *) + using sigma_algebra_Pow[of "space M" "algebra.more M"] + unfolding * . show "finite (space M)" by fact - show "positive \" unfolding positive_def by fact - show "additive M \" unfolding additive_def using assms by simp + show "positive M (measure M)" unfolding positive_def by fact + show "additive M (measure M)" unfolding additive_def using assms by simp qed - then interpret measure_space M \ . + then interpret measure_space M . show "finite_sigma_algebra M" proof show "finite (space M)" by fact @@ -1363,18 +1372,18 @@ qed lemma finite_measure_space_iff: - "finite_measure_space M \ \ - finite (space M) \ sets M = Pow(space M) \ \ (space M) \ \ \ \ {} = 0 \ - (\A\space M. \B\space M. A \ B = {} \ \ (A \ B) = \ A + \ B)" + "finite_measure_space M \ + finite (space M) \ sets M = Pow(space M) \ measure M (space M) \ \ \ measure M {} = 0 \ + (\A\space M. \B\space M. A \ B = {} \ measure M (A \ B) = measure M A + measure M B)" (is "_ = ?rhs") proof (intro iffI) - assume "finite_measure_space M \" - then interpret finite_measure_space M \ . + assume "finite_measure_space M" + then interpret finite_measure_space M . show ?rhs using finite_space sets_eq_Pow measure_additive empty_measure finite_measure by auto next - assume ?rhs then show "finite_measure_space M \" + assume ?rhs then show "finite_measure_space M" by (auto intro!: finite_measure_spaceI) qed diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Probability.thy Wed Feb 02 12:34:45 2011 +0100 @@ -1,6 +1,8 @@ theory Probability imports + Complete_Measure Information "ex/Dining_Cryptographers" + "ex/Koepf_Duermuth_Countermeasure" begin end diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Probability_Space.thy --- a/src/HOL/Probability/Probability_Space.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Probability_Space.thy Wed Feb 02 12:34:45 2011 +0100 @@ -14,7 +14,7 @@ by (cases X) auto locale prob_space = measure_space + - assumes measure_space_1: "\ (space M) = 1" + assumes measure_space_1: "measure M (space M) = 1" lemma abs_real_of_pextreal[simp]: "\real (X :: pextreal)\ = real X" by simp @@ -31,7 +31,7 @@ abbreviation (in prob_space) "prob \ \A. real (\ A)" abbreviation (in prob_space) "prob_preserving \ measure_preserving" abbreviation (in prob_space) "random_variable M' X \ sigma_algebra M' \ X \ measurable M M'" -abbreviation (in prob_space) "expectation \ integral" +abbreviation (in prob_space) "expectation \ integral\<^isup>L M" definition (in prob_space) "indep A B \ A \ events \ B \ events \ prob (A \ B) = prob A * prob B" @@ -193,12 +193,14 @@ lemma (in prob_space) distribution_prob_space: assumes "random_variable S X" - shows "prob_space S (distribution X)" + shows "prob_space (S\measure := distribution X\)" proof - - interpret S: measure_space S "distribution X" unfolding distribution_def - using assms by (intro measure_space_vimage) auto + interpret S: measure_space "S\measure := distribution X\" + unfolding distribution_def using assms + by (intro measure_space_vimage) + (auto intro!: sigma_algebra.sigma_algebra_cong[of S]) show ?thesis - proof + proof (default, simp) have "X -` space S \ space M = space M" using `random_variable S X` by (auto simp: measurable_def) then show "distribution X (space S) = 1" @@ -207,10 +209,10 @@ qed lemma (in prob_space) AE_distribution: - assumes X: "random_variable MX X" and "measure_space.almost_everywhere MX (distribution X) (\x. Q x)" + assumes X: "random_variable MX X" and "measure_space.almost_everywhere (MX\measure := distribution X\) (\x. Q x)" shows "AE x. Q (X x)" proof - - interpret X: prob_space MX "distribution X" using X by (rule distribution_prob_space) + interpret X: prob_space "MX\measure := distribution X\" using X by (rule distribution_prob_space) obtain N where N: "N \ sets MX" "distribution X N = 0" "{x\space MX. \ Q x} \ N" using assms unfolding X.almost_everywhere_def by auto show "AE x. Q (X x)" @@ -228,12 +230,10 @@ lemma (in prob_space) distribution_lebesgue_thm2: assumes "random_variable S X" and "A \ sets S" - shows "distribution X A = - measure_space.positive_integral S (distribution X) (indicator A)" - (is "_ = measure_space.positive_integral _ ?D _") + shows "distribution X A = integral\<^isup>P (S\measure := distribution X\) (indicator A)" proof - - interpret S: prob_space S "distribution X" using assms(1) by (rule distribution_prob_space) - + interpret S: prob_space "S\measure := distribution X\" + using assms(1) by (rule distribution_prob_space) show ?thesis using S.positive_integral_indicator(1) using assms unfolding distribution_def by auto @@ -249,7 +249,7 @@ qed lemma (in prob_space) finite_expectation: - assumes "finite (space M)" "random_variable borel X" + assumes "finite (X`space M)" "random_variable borel X" shows "expectation X = (\ r \ X ` (space M). r * real (distribution X {r}))" using assms unfolding distribution_def using finite_expectation1 by auto @@ -290,27 +290,24 @@ assumes "distribution X {x} = 1" assumes "y \ x" shows "distribution X {y} = 0" -proof - - from distribution_prob_space[OF X] - interpret S: prob_space ?S "distribution X" by simp - have x: "{x} \ sets ?S" - proof (rule ccontr) - assume "{x} \ sets ?S" - hence "X -` {x} \ space M = {}" by auto - thus "False" using assms unfolding distribution_def by auto - qed - have [simp]: "{y} \ {x} = {}" "{x} - {y} = {x}" using `y \ x` by auto - show ?thesis - proof cases - assume "{y} \ sets ?S" - with `{x} \ sets ?S` assms show "distribution X {y} = 0" - using S.measure_inter_full_set[of "{y}" "{x}"] - by simp - next - assume "{y} \ sets ?S" - hence "X -` {y} \ space M = {}" by auto - thus "distribution X {y} = 0" unfolding distribution_def by auto - qed +proof cases + { fix x have "X -` {x} \ space M \ sets M" + proof cases + assume "x \ X`space M" with X show ?thesis + by (auto simp: measurable_def image_iff) + next + assume "x \ X`space M" then have "X -` {x} \ space M = {}" by auto + then show ?thesis by auto + qed } note single = this + have "X -` {x} \ space M - X -` {y} \ space M = X -` {x} \ space M" + "X -` {y} \ space M \ (X -` {x} \ space M) = {}" + using `y \ x` by auto + with measure_inter_full_set[OF single single, of x y] assms(2) + show ?thesis unfolding distribution_def measure_space_1 by auto +next + assume "{y} \ sets ?S" + then have "X -` {y} \ space M = {}" by auto + thus "distribution X {y} = 0" unfolding distribution_def by auto qed lemma (in prob_space) joint_distribution_Times_le_fst: @@ -344,14 +341,14 @@ lemma (in prob_space) random_variable_pairI: assumes "random_variable MX X" assumes "random_variable MY Y" - shows "random_variable (sigma (pair_algebra MX MY)) (\x. (X x, Y x))" + shows "random_variable (MX \\<^isub>M MY) (\x. (X x, Y x))" proof interpret MX: sigma_algebra MX using assms by simp interpret MY: sigma_algebra MY using assms by simp interpret P: pair_sigma_algebra MX MY by default - show "sigma_algebra (sigma (pair_algebra MX MY))" by default + show "sigma_algebra (MX \\<^isub>M MY)" by default have sa: "sigma_algebra M" by default - show "(\x. (X x, Y x)) \ measurable M (sigma (pair_algebra MX MY))" + show "(\x. (X x, Y x)) \ measurable M (MX \\<^isub>M MY)" unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) qed @@ -377,15 +374,12 @@ joint_distribution (\x. (X x, Y x)) Z {((x, y), z)}" unfolding distribution_def by (auto intro!: arg_cong[where f=\]) -locale pair_prob_space = M1: prob_space M1 p1 + M2: prob_space M2 p2 for M1 p1 M2 p2 - -sublocale pair_prob_space \ pair_sigma_finite M1 p1 M2 p2 by default +locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2 -sublocale pair_prob_space \ P: prob_space P pair_measure -proof - show "pair_measure (space P) = 1" - by (simp add: pair_algebra_def pair_measure_times M1.measure_space_1 M2.measure_space_1) -qed +sublocale pair_prob_space \ pair_sigma_finite M1 M2 by default + +sublocale pair_prob_space \ P: prob_space P +by default (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure) lemma countably_additiveI[case_names countably]: assumes "\A. \ range A \ sets M ; disjoint_family A ; (\i. A i) \ sets M\ \ @@ -395,38 +389,8 @@ lemma (in prob_space) joint_distribution_prob_space: assumes "random_variable MX X" "random_variable MY Y" - shows "prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)" -proof - - interpret X: prob_space MX "distribution X" by (intro distribution_prob_space assms) - interpret Y: prob_space MY "distribution Y" by (intro distribution_prob_space assms) - interpret XY: pair_sigma_finite MX "distribution X" MY "distribution Y" by default - show ?thesis - proof - let "?X A" = "(\x. (X x, Y x)) -` A \ space M" - show "joint_distribution X Y {} = 0" by (simp add: distribution_def) - show "countably_additive XY.P (joint_distribution X Y)" - proof (rule countably_additiveI) - fix A :: "nat \ ('b \ 'c) set" - assume A: "range A \ sets XY.P" and df: "disjoint_family A" - have "(\\<^isub>\n. \ (?X (A n))) = \ (\x. ?X (A x))" - proof (intro measure_countably_additive) - have "sigma_algebra M" by default - then have *: "(\x. (X x, Y x)) \ measurable M XY.P" - using assms by (simp add: XY.measurable_pair comp_def) - show "range (\n. ?X (A n)) \ events" - using measurable_sets[OF *] A by auto - show "disjoint_family (\n. ?X (A n))" - by (intro disjoint_family_on_bisimulation[OF df]) auto - qed - then show "(\\<^isub>\n. joint_distribution X Y (A n)) = joint_distribution X Y (\i. A i)" - by (simp add: distribution_def vimage_UN) - qed - have "?X (space MX \ space MY) = space M" - using assms by (auto simp: measurable_def) - then show "joint_distribution X Y (space XY.P) = 1" - by (simp add: space_pair_algebra distribution_def measure_space_1) - qed -qed + shows "prob_space ((MX \\<^isub>M MY) \ measure := joint_distribution X Y\)" + using random_variable_pairI[OF assms] by (rule distribution_prob_space) section "Probability spaces on finite sets" @@ -443,32 +407,32 @@ lemma (in prob_space) distribution_finite_prob_space: assumes "finite_random_variable MX X" - shows "finite_prob_space MX (distribution X)" + shows "finite_prob_space (MX\measure := distribution X\)" proof - - interpret X: prob_space MX "distribution X" + interpret X: prob_space "MX\measure := distribution X\" using assms[THEN finite_random_variableD] by (rule distribution_prob_space) interpret MX: finite_sigma_algebra MX - using assms by simp + using assms by auto show ?thesis - proof + proof (default, simp_all) fix x assume "x \ space MX" then have "X -` {x} \ space M \ sets M" using assms unfolding measurable_def by simp then show "distribution X {x} \ \" unfolding distribution_def by simp - qed + qed (rule MX.finite_space) qed lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]: - assumes "simple_function X" - shows "finite_random_variable \ space = X`space M, sets = Pow (X`space M) \ X" + assumes "simple_function M X" + shows "finite_random_variable \ space = X`space M, sets = Pow (X`space M), \ = x \ X" + (is "finite_random_variable ?X _") proof (intro conjI) have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp - interpret X: sigma_algebra "\space = X ` space M, sets = Pow (X ` space M)\" - by (rule sigma_algebra_Pow) - show "finite_sigma_algebra \space = X ` space M, sets = Pow (X ` space M)\" + interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow) + show "finite_sigma_algebra ?X" by default auto - show "X \ measurable M \space = X ` space M, sets = Pow (X ` space M)\" + show "X \ measurable M ?X" proof (unfold measurable_def, clarsimp) fix A assume A: "A \ X`space M" then have "finite A" by (rule finite_subset) simp @@ -481,13 +445,13 @@ qed lemma (in prob_space) simple_function_imp_random_variable[simp, intro]: - assumes "simple_function X" - shows "random_variable \ space = X`space M, sets = Pow (X`space M) \ X" - using simple_function_imp_finite_random_variable[OF assms] + assumes "simple_function M X" + shows "random_variable \ space = X`space M, sets = Pow (X`space M), \ = ext \ X" + using simple_function_imp_finite_random_variable[OF assms, of ext] by (auto dest!: finite_random_variableD) lemma (in prob_space) sum_over_space_real_distribution: - "simple_function X \ (\x\X`space M. real (distribution X {x})) = 1" + "simple_function M X \ (\x\X`space M. real (distribution X {x})) = 1" unfolding distribution_def prob_space[symmetric] by (subst real_finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def simple_function_def @@ -496,14 +460,14 @@ lemma (in prob_space) finite_random_variable_pairI: assumes "finite_random_variable MX X" assumes "finite_random_variable MY Y" - shows "finite_random_variable (sigma (pair_algebra MX MY)) (\x. (X x, Y x))" + shows "finite_random_variable (MX \\<^isub>M MY) (\x. (X x, Y x))" proof interpret MX: finite_sigma_algebra MX using assms by simp interpret MY: finite_sigma_algebra MY using assms by simp interpret P: pair_finite_sigma_algebra MX MY by default - show "finite_sigma_algebra (sigma (pair_algebra MX MY))" by default + show "finite_sigma_algebra (MX \\<^isub>M MY)" by default have sa: "sigma_algebra M" by default - show "(\x. (X x, Y x)) \ measurable M (sigma (pair_algebra MX MY))" + show "(\x. (X x, Y x)) \ measurable M (MX \\<^isub>M MY)" unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) qed @@ -599,6 +563,7 @@ finite_random_variable_imp_sets[OF Y]] by simp lemma (in prob_space) setsum_real_joint_distribution: + fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme" assumes X: "finite_random_variable MX X" assumes Y: "random_variable MY Y" "B \ sets MY" shows "(\a\space MX. real (joint_distribution X Y ({a} \ B))) = real (distribution Y B)" @@ -607,10 +572,11 @@ show ?thesis unfolding setsum_joint_distribution[OF assms, symmetric] using distribution_finite[OF random_variable_pairI[OF finite_random_variableD[OF X] Y(1)]] Y(2) - by (simp add: space_pair_algebra in_sigma pair_algebraI MX.sets_eq_Pow real_of_pextreal_setsum) + by (simp add: space_pair_measure real_of_pextreal_setsum) qed lemma (in prob_space) setsum_real_joint_distribution_singleton: + fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme" assumes X: "finite_random_variable MX X" assumes Y: "finite_random_variable MY Y" "b \ space MY" shows "(\a\space MX. real (joint_distribution X Y {(a,b)})) = real (distribution Y {b})" @@ -618,39 +584,20 @@ finite_random_variableD[OF Y(1)] finite_random_variable_imp_sets[OF Y]] by simp -locale pair_finite_prob_space = M1: finite_prob_space M1 p1 + M2: finite_prob_space M2 p2 for M1 p1 M2 p2 +locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 -sublocale pair_finite_prob_space \ pair_prob_space M1 p1 M2 p2 by default -sublocale pair_finite_prob_space \ pair_finite_space M1 p1 M2 p2 by default -sublocale pair_finite_prob_space \ finite_prob_space P pair_measure by default +sublocale pair_finite_prob_space \ pair_prob_space M1 M2 by default +sublocale pair_finite_prob_space \ pair_finite_space M1 M2 by default +sublocale pair_finite_prob_space \ finite_prob_space P by default lemma (in prob_space) joint_distribution_finite_prob_space: assumes X: "finite_random_variable MX X" assumes Y: "finite_random_variable MY Y" - shows "finite_prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)" -proof - - interpret X: finite_prob_space MX "distribution X" - using X by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space MY "distribution Y" - using Y by (rule distribution_finite_prob_space) - interpret P: prob_space "sigma (pair_algebra MX MY)" "joint_distribution X Y" - using assms[THEN finite_random_variableD] by (rule joint_distribution_prob_space) - interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y" - by default - show ?thesis - proof - fix x assume "x \ space XY.P" - moreover have "(\x. (X x, Y x)) \ measurable M XY.P" - using X Y by (intro XY.measurable_pair) (simp_all add: o_def, default) - ultimately have "(\x. (X x, Y x)) -` {x} \ space M \ sets M" - unfolding measurable_def by simp - then show "joint_distribution X Y {x} \ \" - unfolding distribution_def by simp - qed -qed + shows "finite_prob_space ((MX \\<^isub>M MY)\ measure := joint_distribution X Y\)" + by (intro distribution_finite_prob_space finite_random_variable_pairI X Y) lemma finite_prob_space_eq: - "finite_prob_space M \ \ finite_measure_space M \ \ \ (space M) = 1" + "finite_prob_space M \ finite_measure_space M \ measure M (space M) = 1" unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def by auto @@ -852,48 +799,51 @@ lemma (in prob_space) prob_space_subalgebra: assumes "sigma_algebra N" "sets N \ sets M" "space N = space M" - shows "prob_space N \" + and "\A. A \ sets N \ measure N A = \ A" + shows "prob_space N" proof - - interpret N: measure_space N \ - using measure_space_subalgebra[OF assms] . + interpret N: measure_space N + by (rule measure_space_subalgebra[OF assms]) show ?thesis - proof qed (simp add: `space N = space M` measure_space_1) + proof qed (insert assms(4)[OF N.top], simp add: assms measure_space_1) qed lemma (in prob_space) prob_space_of_restricted_space: assumes "\ A \ 0" "\ A \ \" "A \ sets M" - shows "prob_space (restricted_space A) (\S. \ S / \ A)" - unfolding prob_space_def prob_space_axioms_def -proof - show "\ (space (restricted_space A)) / \ A = 1" - using `\ A \ 0` `\ A \ \` by (auto simp: pextreal_noteq_omega_Ex) - have *: "\S. \ S / \ A = inverse (\ A) * \ S" by (simp add: mult_commute) - interpret A: measure_space "restricted_space A" \ + shows "prob_space (restricted_space A \measure := \S. \ S / \ A\)" + (is "prob_space ?P") +proof - + interpret A: measure_space "restricted_space A" using `A \ sets M` by (rule restricted_measure_space) - show "measure_space (restricted_space A) (\S. \ S / \ A)" + interpret A': sigma_algebra ?P + by (rule A.sigma_algebra_cong) auto + show "prob_space ?P" proof - show "\ {} / \ A = 0" by auto - show "countably_additive (restricted_space A) (\S. \ S / \ A)" - unfolding countably_additive_def psuminf_cmult_right * + show "measure ?P (space ?P) = 1" + using `\ A \ 0` `\ A \ \` by (auto simp: pextreal_noteq_omega_Ex) + show "measure ?P {} = 0" by auto + have "\S. \ S / \ A = inverse (\ A) * \ S" by (simp add: mult_commute) + then show "countably_additive ?P (measure ?P)" + unfolding countably_additive_def psuminf_cmult_right using A.measure_countably_additive by auto qed qed lemma finite_prob_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" "\ (space M) = 1" "\ {} = 0" - and "\A B. A\space M \ B\space M \ A \ B = {} \ \ (A \ B) = \ A + \ B" - shows "finite_prob_space M \" + assumes "finite (space M)" "sets M = Pow(space M)" "measure M (space M) = 1" "measure M {} = 0" + and "\A B. A\space M \ B\space M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" + shows "finite_prob_space M" unfolding finite_prob_space_eq proof - show "finite_measure_space M \" using assms + show "finite_measure_space M" using assms by (auto intro!: finite_measure_spaceI) - show "\ (space M) = 1" by fact + show "measure M (space M) = 1" by fact qed lemma (in finite_prob_space) finite_measure_space: fixes X :: "'a \ 'x" - shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" - (is "finite_measure_space ?S _") + shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" + (is "finite_measure_space ?S") proof (rule finite_measure_spaceI, simp_all) show "finite (X ` space M)" using finite_space by simp next @@ -905,7 +855,7 @@ qed lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" + "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M), measure = distribution X \" by (simp add: finite_prob_space_eq finite_measure_space) lemma (in finite_prob_space) real_distribution_order': @@ -919,8 +869,8 @@ lemma (in finite_prob_space) finite_product_measure_space: fixes X :: "'a \ 'x" and Y :: "'a \ 'y" assumes "finite s1" "finite s2" - shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2)\ (joint_distribution X Y)" - (is "finite_measure_space ?M ?D") + shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2), measure = joint_distribution X Y\" + (is "finite_measure_space ?M") proof (rule finite_measure_spaceI, simp_all) show "finite (s1 \ s2)" using assms by auto @@ -936,14 +886,14 @@ lemma (in finite_prob_space) finite_product_measure_space_of_images: shows "finite_measure_space \ space = X ` space M \ Y ` space M, - sets = Pow (X ` space M \ Y ` space M) \ - (joint_distribution X Y)" + sets = Pow (X ` space M \ Y ` space M), + measure = joint_distribution X Y \" using finite_space by (auto intro!: finite_product_measure_space) lemma (in finite_prob_space) finite_product_prob_space_of_images: - "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M)\ - (joint_distribution X Y)" - (is "finite_prob_space ?S _") + "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M), + measure = joint_distribution X Y \" + (is "finite_prob_space ?S") proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) have "X -` X ` space M \ Y -` Y ` space M \ space M = space M" by auto thus "joint_distribution X Y (X ` space M \ Y ` space M) = 1" @@ -953,27 +903,29 @@ section "Conditional Expectation and Probability" lemma (in prob_space) conditional_expectation_exists: - fixes X :: "'a \ pextreal" + fixes X :: "'a \ pextreal" and N :: "('a, 'b) measure_space_scheme" assumes borel: "X \ borel_measurable M" - and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" + and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" "\A. measure N A = \ A" shows "\Y\borel_measurable N. \C\sets N. - (\\<^isup>+x. Y x * indicator C x) = (\\<^isup>+x. X x * indicator C x)" + (\\<^isup>+x. Y x * indicator C x \M) = (\\<^isup>+x. X x * indicator C x \M)" proof - - interpret P: prob_space N \ + note N(4)[simp] + interpret P: prob_space N using prob_space_subalgebra[OF N] . let "?f A" = "\x. X x * indicator A x" - let "?Q A" = "positive_integral (?f A)" + let "?Q A" = "integral\<^isup>P M (?f A)" from measure_space_density[OF borel] - have Q: "measure_space N ?Q" - by (rule measure_space.measure_space_subalgebra[OF _ N]) - then interpret Q: measure_space N ?Q . + have Q: "measure_space (N\ measure := ?Q \)" + apply (rule measure_space.measure_space_subalgebra[of "M\ measure := ?Q \"]) + using N by (auto intro!: P.sigma_algebra_cong) + then interpret Q: measure_space "N\ measure := ?Q \" . have "P.absolutely_continuous ?Q" unfolding P.absolutely_continuous_def proof safe - fix A assume "A \ sets N" "\ A = 0" + fix A assume "A \ sets N" "P.\ A = 0" moreover then have f_borel: "?f A \ borel_measurable M" using borel N by (auto intro: borel_measurable_indicator) moreover have "{x\space M. ?f A x \ 0} = (?f A -` {0<..} \ space M) \ A" @@ -986,28 +938,28 @@ qed from P.Radon_Nikodym[OF Q this] obtain Y where Y: "Y \ borel_measurable N" - "\A. A \ sets N \ ?Q A = P.positive_integral (\x. Y x * indicator A x)" + "\A. A \ sets N \ ?Q A =(\\<^isup>+x. Y x * indicator A x \N)" by blast with N(2) show ?thesis - by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,1)]) + by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,4,1)]) qed definition (in prob_space) "conditional_expectation N X = (SOME Y. Y\borel_measurable N - \ (\C\sets N. (\\<^isup>+x. Y x * indicator C x) = (\\<^isup>+x. X x * indicator C x)))" + \ (\C\sets N. (\\<^isup>+x. Y x * indicator C x\M) = (\\<^isup>+x. X x * indicator C x\M)))" abbreviation (in prob_space) "conditional_prob N A \ conditional_expectation N (indicator A)" lemma (in prob_space) - fixes X :: "'a \ pextreal" + fixes X :: "'a \ pextreal" and N :: "('a, 'b) measure_space_scheme" assumes borel: "X \ borel_measurable M" - and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" + and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" "\A. measure N A = \ A" shows borel_measurable_conditional_expectation: "conditional_expectation N X \ borel_measurable N" and conditional_expectation: "\C. C \ sets N \ - (\\<^isup>+x. conditional_expectation N X x * indicator C x) = - (\\<^isup>+x. X x * indicator C x)" + (\\<^isup>+x. conditional_expectation N X x * indicator C x \M) = + (\\<^isup>+x. X x * indicator C x \M)" (is "\C. C \ sets N \ ?eq C") proof - note CE = conditional_expectation_exists[OF assms, unfolded Bex_def] @@ -1042,21 +994,21 @@ assume "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" from va.borel_measurable_implies_simple_function_sequence[OF this] - obtain f where f: "\i. va.simple_function (f i)" and "f \ Z" by blast + obtain f where f: "\i. simple_function (M'.vimage_algebra (space M) Y) (f i)" and "f \ Z" by blast - have "\i. \g. M'.simple_function g \ (\x\space M. f i x = g (Y x))" + have "\i. \g. simple_function M' g \ (\x\space M. f i x = g (Y x))" proof fix i from f[of i] have "finite (f i`space M)" and B_ex: "\z\(f i)`space M. \B. B \ sets M' \ (f i) -` {z} \ space M = Y -` B \ space M" - unfolding va.simple_function_def by auto + unfolding simple_function_def by auto from B_ex[THEN bchoice] guess B .. note B = this let ?g = "\x. \z\f i`space M. z * indicator (B z) x" - show "\g. M'.simple_function g \ (\x\space M. f i x = g (Y x))" + show "\g. simple_function M' g \ (\x\space M. f i x = g (Y x))" proof (intro exI[of _ ?g] conjI ballI) - show "M'.simple_function ?g" using B by auto + show "simple_function M' ?g" using B by auto fix x assume "x \ space M" then have "\z. z \ f i`space M \ indicator (B z) (Y x) = (indicator (f i -` {z} \ space M) x::pextreal)" diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Product_Measure.thy --- a/src/HOL/Probability/Product_Measure.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Product_Measure.thy Wed Feb 02 12:34:45 2011 +0100 @@ -2,6 +2,20 @@ imports Lebesgue_Integration begin +lemma measurable_cancel_measure2[simp]: + "measurable M1 (M2\measure := m\) = measurable M1 M2" + unfolding measurable_def by auto + +lemma measurable_cancel_measure1[simp]: + "measurable (M1\measure := m\) M2 = measurable M1 M2" + unfolding measurable_def by auto + +lemma sigma_sets_subseteq: assumes "A \ B" shows "sigma_sets X A \ sigma_sets X B" +proof + fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" + by induct (insert `A \ B`, auto intro: sigma_sets.intros) +qed + lemma times_Int_times: "A \ B \ C \ D = (A \ C) \ (B \ D)" by auto @@ -20,6 +34,18 @@ abbreviation "Pi\<^isub>E A B \ Pi A B \ extensional A" +syntax + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10) + +syntax (xsymbols) + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + +syntax (HTML output) + "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + +translations + "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" + abbreviation funcset_extensional :: "['a set, 'b set] => ('a => 'b) set" (infixr "->\<^isub>E" 60) where @@ -171,61 +197,121 @@ "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" by (auto simp: fun_eq_iff) +lemma Pi_eq_subset: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \ I" + shows "F i \ F' i" +proof + fix x assume "x \ F i" + with ne have "\j. \y. ((j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined))" by auto + from choice[OF this] guess f .. note f = this + then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) + then have "f \ Pi\<^isub>E I F'" using assms by simp + then show "x \ F' i" using f `i \ I` by auto +qed + +lemma Pi_eq_iff_not_empty: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i)" +proof (intro iffI ballI) + fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \ I" + show "F i = F' i" + using Pi_eq_subset[of I F F', OF ne eq i] + using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] + by auto +qed auto + +lemma Pi_eq_empty_iff: + "Pi\<^isub>E I F = {} \ (\i\I. F i = {})" +proof + assume "Pi\<^isub>E I F = {}" + show "\i\I. F i = {}" + proof (rule ccontr) + assume "\ ?thesis" + then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" by auto + from choice[OF this] guess f .. + then have "f \ Pi\<^isub>E I F" by (auto simp: extensional_def) + with `Pi\<^isub>E I F = {}` show False by auto + qed +qed auto + +lemma Pi_eq_iff: + "Pi\<^isub>E I F = Pi\<^isub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" +proof (intro iffI disjCI) + assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'" + assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" + then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" + using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto + with Pi_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" by auto +next + assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" + then show "Pi\<^isub>E I F = Pi\<^isub>E I F'" + using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto +qed + section "Binary products" definition - "pair_algebra A B = \ space = space A \ space B, - sets = {a \ b | a b. a \ sets A \ b \ sets B} \" + "pair_measure_generator A B = + \ space = space A \ space B, + sets = {a \ b | a b. a \ sets A \ b \ sets B}, + measure = \X. \\<^isup>+x. (\\<^isup>+y. indicator X (x,y) \B) \A \" + +definition pair_measure (infixr "\\<^isub>M" 80) where + "A \\<^isub>M B = sigma (pair_measure_generator A B)" locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2 - for M1 M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + +abbreviation (in pair_sigma_algebra) + "E \ pair_measure_generator M1 M2" abbreviation (in pair_sigma_algebra) - "E \ pair_algebra M1 M2" + "P \ M1 \\<^isub>M M2" -abbreviation (in pair_sigma_algebra) - "P \ sigma E" +lemma sigma_algebra_pair_measure: + "sets M1 \ Pow (space M1) \ sets M2 \ Pow (space M2) \ sigma_algebra (pair_measure M1 M2)" + by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma) sublocale pair_sigma_algebra \ sigma_algebra P - using M1.sets_into_space M2.sets_into_space - by (force simp: pair_algebra_def intro!: sigma_algebra_sigma) + using M1.space_closed M2.space_closed + by (rule sigma_algebra_pair_measure) -lemma pair_algebraI[intro, simp]: - "x \ sets A \ y \ sets B \ x \ y \ sets (pair_algebra A B)" - by (auto simp add: pair_algebra_def) +lemma pair_measure_generatorI[intro, simp]: + "x \ sets A \ y \ sets B \ x \ y \ sets (pair_measure_generator A B)" + by (auto simp add: pair_measure_generator_def) -lemma space_pair_algebra: - "space (pair_algebra A B) = space A \ space B" - by (simp add: pair_algebra_def) +lemma pair_measureI[intro, simp]: + "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" + by (auto simp add: pair_measure_def) -lemma sets_pair_algebra: "sets (pair_algebra N M) = (\(x, y). x \ y) ` (sets N \ sets M)" - unfolding pair_algebra_def by auto +lemma space_pair_measure: + "space (A \\<^isub>M B) = space A \ space B" + by (simp add: pair_measure_def pair_measure_generator_def) -lemma pair_algebra_sets_into_space: - assumes "sets M \ Pow (space M)" "sets N \ Pow (space N)" - shows "sets (pair_algebra M N) \ Pow (space (pair_algebra M N))" - using assms by (auto simp: pair_algebra_def) +lemma sets_pair_measure_generator: + "sets (pair_measure_generator N M) = (\(x, y). x \ y) ` (sets N \ sets M)" + unfolding pair_measure_generator_def by auto -lemma pair_algebra_Int_snd: +lemma pair_measure_generator_sets_into_space: + assumes "sets M \ Pow (space M)" "sets N \ Pow (space N)" + shows "sets (pair_measure_generator M N) \ Pow (space (pair_measure_generator M N))" + using assms by (auto simp: pair_measure_generator_def) + +lemma pair_measure_generator_Int_snd: assumes "sets S1 \ Pow (space S1)" - shows "pair_algebra S1 (algebra.restricted_space S2 A) = - algebra.restricted_space (pair_algebra S1 S2) (space S1 \ A)" + shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) = + sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \ A))" (is "?L = ?R") -proof (intro algebra.equality set_eqI iffI) - fix X assume "X \ sets ?L" - then obtain A1 A2 where X: "X = A1 \ (A \ A2)" and "A1 \ sets S1" "A2 \ sets S2" - by (auto simp: pair_algebra_def) - then show "X \ sets ?R" unfolding pair_algebra_def - using assms apply simp by (intro image_eqI[of _ _ "A1 \ A2"]) auto -next - fix X assume "X \ sets ?R" - then obtain A1 A2 where "X = space S1 \ A \ A1 \ A2" "A1 \ sets S1" "A2 \ sets S2" - by (auto simp: pair_algebra_def) - moreover then have "X = A1 \ (A \ A2)" - using assms by auto - ultimately show "X \ sets ?L" - unfolding pair_algebra_def by auto -qed (auto simp add: pair_algebra_def) + apply (auto simp: pair_measure_generator_def image_iff) + using assms + apply (rule_tac x="a \ xa" in exI) + apply force + using assms + apply (rule_tac x="a" in exI) + apply (rule_tac x="b \ A" in exI) + apply auto + done lemma (in pair_sigma_algebra) shows measurable_fst[intro!, simp]: @@ -243,7 +329,7 @@ apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) using M2.sets_into_space by force+ } ultimately have "?fst \ ?snd" - by (fastsimp simp: measurable_def sets_sigma space_pair_algebra + by (fastsimp simp: measurable_def sets_sigma space_pair_measure intro!: sigma_sets.Basic) then show ?fst ?snd by auto qed @@ -257,15 +343,15 @@ from assms show ?thesis proof (safe intro!: measurable_comp[where b=P]) assume f: "(fst \ f) \ measurable M M1" and s: "(snd \ f) \ measurable M M2" - show "f \ measurable M P" + show "f \ measurable M P" unfolding pair_measure_def proof (rule M.measurable_sigma) - show "sets (pair_algebra M1 M2) \ Pow (space E)" - unfolding pair_algebra_def using M1.sets_into_space M2.sets_into_space by auto + show "sets (pair_measure_generator M1 M2) \ Pow (space E)" + unfolding pair_measure_generator_def using M1.sets_into_space M2.sets_into_space by auto show "f \ space M \ space E" - using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma space_pair_algebra) + using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma pair_measure_generator_def) fix A assume "A \ sets E" then obtain B C where "B \ sets M1" "C \ sets M2" "A = B \ C" - unfolding pair_algebra_def by auto + unfolding pair_measure_generator_def by auto moreover have "(fst \ f) -` B \ space M \ sets M" using f `B \ sets M1` unfolding measurable_def by auto moreover have "(snd \ f) -` C \ space M \ sets M" @@ -283,52 +369,58 @@ shows "f \ measurable M P" unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp -lemma pair_algebraE: - assumes "X \ sets (pair_algebra M1 M2)" +lemma pair_measure_generatorE: + assumes "X \ sets (pair_measure_generator M1 M2)" obtains A B where "X = A \ B" "A \ sets M1" "B \ sets M2" - using assms unfolding pair_algebra_def by auto + using assms unfolding pair_measure_generator_def by auto -lemma (in pair_sigma_algebra) pair_algebra_swap: - "(\X. (\(x,y). (y,x)) -` X \ space M2 \ space M1) ` sets E = sets (pair_algebra M2 M1)" -proof (safe elim!: pair_algebraE) +lemma (in pair_sigma_algebra) pair_measure_generator_swap: + "(\X. (\(x,y). (y,x)) -` X \ space M2 \ space M1) ` sets E = sets (pair_measure_generator M2 M1)" +proof (safe elim!: pair_measure_generatorE) fix A B assume "A \ sets M1" "B \ sets M2" moreover then have "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 = B \ A" using M1.sets_into_space M2.sets_into_space by auto - ultimately show "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 \ sets (pair_algebra M2 M1)" - by (auto intro: pair_algebraI) + ultimately show "(\(x, y). (y, x)) -` (A \ B) \ space M2 \ space M1 \ sets (pair_measure_generator M2 M1)" + by (auto intro: pair_measure_generatorI) next fix A B assume "A \ sets M1" "B \ sets M2" then show "B \ A \ (\X. (\(x, y). (y, x)) -` X \ space M2 \ space M1) ` sets E" using M1.sets_into_space M2.sets_into_space - by (auto intro!: image_eqI[where x="A \ B"] pair_algebraI) + by (auto intro!: image_eqI[where x="A \ B"] pair_measure_generatorI) qed lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap: assumes Q: "Q \ sets P" - shows "(\(x,y). (y, x)) ` Q \ sets (sigma (pair_algebra M2 M1))" (is "_ \ sets ?Q") + shows "(\(x,y). (y, x)) -` Q \ sets (M2 \\<^isub>M M1)" (is "_ \ sets ?Q") proof - - have *: "(\(x,y). (y, x)) \ space M2 \ space M1 \ (space M1 \ space M2)" - "sets (pair_algebra M1 M2) \ Pow (space M1 \ space M2)" - using M1.sets_into_space M2.sets_into_space by (auto elim!: pair_algebraE) - from Q sets_into_space show ?thesis - by (auto intro!: image_eqI[where x=Q] - simp: pair_algebra_swap[symmetric] sets_sigma - sigma_sets_vimage[OF *] space_pair_algebra) + let "?f Q" = "(\(x,y). (y, x)) -` Q \ space M2 \ space M1" + have *: "(\(x,y). (y, x)) -` Q = ?f Q" + using sets_into_space[OF Q] by (auto simp: space_pair_measure) + have "sets (M2 \\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))" + unfolding pair_measure_def .. + also have "\ = sigma_sets (space M2 \ space M1) (?f ` sets E)" + unfolding sigma_def pair_measure_generator_swap[symmetric] + by (simp add: pair_measure_generator_def) + also have "\ = ?f ` sigma_sets (space M1 \ space M2) (sets E)" + using M1.sets_into_space M2.sets_into_space + by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def) + also have "\ = ?f ` sets P" + unfolding pair_measure_def pair_measure_generator_def sigma_def by simp + finally show ?thesis + using Q by (subst *) auto qed lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable: - shows "(\(x,y). (y, x)) \ measurable P (sigma (pair_algebra M2 M1))" + shows "(\(x,y). (y, x)) \ measurable P (M2 \\<^isub>M M1)" (is "?f \ measurable ?P ?Q") unfolding measurable_def proof (intro CollectI conjI Pi_I ballI) fix x assume "x \ space ?P" then show "(case x of (x, y) \ (y, x)) \ space ?Q" - unfolding pair_algebra_def by auto + unfolding pair_measure_generator_def pair_measure_def by auto next - fix A assume "A \ sets ?Q" + fix A assume "A \ sets (M2 \\<^isub>M M1)" interpret Q: pair_sigma_algebra M2 M1 by default - have "?f -` A \ space ?P = (\(x,y). (y, x)) ` A" - using Q.sets_into_space `A \ sets ?Q` by (auto simp: pair_algebra_def) - with Q.sets_pair_sigma_algebra_swap[OF `A \ sets ?Q`] + with Q.sets_pair_sigma_algebra_swap[OF `A \ sets (M2 \\<^isub>M M1)`] show "?f -` A \ space ?P \ sets ?P" by simp qed @@ -338,13 +430,13 @@ let ?Q' = "{Q. Q \ space P \ Pair x -` Q \ sets M2}" let ?Q = "\ space = space P, sets = ?Q' \" interpret Q: sigma_algebra ?Q - proof qed (auto simp: vimage_UN vimage_Diff space_pair_algebra) + proof qed (auto simp: vimage_UN vimage_Diff space_pair_measure) have "sets E \ sets ?Q" using M1.sets_into_space M2.sets_into_space - by (auto simp: pair_algebra_def space_pair_algebra) + by (auto simp: pair_measure_generator_def space_pair_measure) then have "sets P \ sets ?Q" - by (subst pair_algebra_def, intro Q.sets_sigma_subset) - (simp_all add: pair_algebra_def) + apply (subst pair_measure_def, intro Q.sets_sigma_subset) + by (simp add: pair_measure_def) with assms show ?thesis by auto qed @@ -352,9 +444,8 @@ assumes Q: "Q \ sets P" shows "(\x. (x, y)) -` Q \ sets M1" (is "?cut Q \ sets M1") proof - interpret Q: pair_sigma_algebra M2 M1 by default - have "Pair y -` (\(x, y). (y, x)) ` Q = (\x. (x, y)) -` Q" by auto with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y] - show ?thesis by simp + show ?thesis by (simp add: vimage_compose[symmetric] comp_def) qed lemma (in pair_sigma_algebra) measurable_pair_image_snd: @@ -363,14 +454,15 @@ unfolding measurable_def proof (intro CollectI conjI Pi_I ballI) fix y assume "y \ space M2" with `f \ measurable P M` `x \ space M1` - show "f (x, y) \ space M" unfolding measurable_def pair_algebra_def by auto + show "f (x, y) \ space M" + unfolding measurable_def pair_measure_generator_def pair_measure_def by auto next fix A assume "A \ sets M" then have "Pair x -` (f -` A \ space P) \ sets M2" (is "?C \ _") using `f \ measurable P M` by (intro measurable_cut_fst) (auto simp: measurable_def) also have "?C = (\y. f (x, y)) -` A \ space M2" - using `x \ space M1` by (auto simp: pair_algebra_def) + using `x \ space M1` by (auto simp: pair_measure_generator_def pair_measure_def) finally show "(\y. f (x, y)) -` A \ space M2 \ sets M2" . qed @@ -384,31 +476,31 @@ show ?thesis by simp qed -lemma (in pair_sigma_algebra) Int_stable_pair_algebra: "Int_stable E" +lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E" unfolding Int_stable_def proof (intro ballI) fix A B assume "A \ sets E" "B \ sets E" then obtain A1 A2 B1 B2 where "A = A1 \ A2" "B = B1 \ B2" "A1 \ sets M1" "A2 \ sets M2" "B1 \ sets M1" "B2 \ sets M2" - unfolding pair_algebra_def by auto + unfolding pair_measure_generator_def by auto then show "A \ B \ sets E" - by (auto simp add: times_Int_times pair_algebra_def) + by (auto simp add: times_Int_times pair_measure_generator_def) qed lemma finite_measure_cut_measurable: - fixes M1 :: "'a algebra" and M2 :: "'b algebra" - assumes "sigma_finite_measure M1 \1" "finite_measure M2 \2" - assumes "Q \ sets (sigma (pair_algebra M1 M2))" - shows "(\x. \2 (Pair x -` Q)) \ borel_measurable M1" + fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" + assumes "sigma_finite_measure M1" "finite_measure M2" + assumes "Q \ sets (M1 \\<^isub>M M2)" + shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") proof - - interpret M1: sigma_finite_measure M1 \1 by fact - interpret M2: finite_measure M2 \2 by fact + interpret M1: sigma_finite_measure M1 by fact + interpret M2: finite_measure M2 by fact interpret pair_sigma_algebra M1 M2 by default have [intro]: "sigma_algebra M1" by fact have [intro]: "sigma_algebra M2" by fact let ?D = "\ space = space P, sets = {A\sets P. ?s A \ borel_measurable M1} \" - note space_pair_algebra[simp] + note space_pair_measure[simp] interpret dynkin_system ?D proof (intro dynkin_systemI) fix A assume "A \ sets ?D" then show "A \ space ?D" @@ -418,73 +510,83 @@ by (auto simp add: if_distrib intro!: M1.measurable_If) next fix A assume "A \ sets ?D" - with sets_into_space have "\x. \2 (Pair x -` (space M1 \ space M2 - A)) = - (if x \ space M1 then \2 (space M2) - ?s A x else 0)" + with sets_into_space have "\x. measure M2 (Pair x -` (space M1 \ space M2 - A)) = + (if x \ space M1 then measure M2 (space M2) - ?s A x else 0)" by (auto intro!: M2.finite_measure_compl measurable_cut_fst simp: vimage_Diff) with `A \ sets ?D` top show "space ?D - A \ sets ?D" by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pextreal_diff) next fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ sets ?D" - moreover then have "\x. \2 (\i. Pair x -` F i) = (\\<^isub>\ i. ?s (F i) x)" + moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\\<^isub>\ i. ?s (F i) x)" by (intro M2.measure_countably_additive[symmetric]) (auto intro!: measurable_cut_fst simp: disjoint_family_on_def) ultimately show "(\i. F i) \ sets ?D" by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf) qed - have "P = ?D" + have "sets P = sets ?D" apply (subst pair_measure_def) proof (intro dynkin_lemma) - show "Int_stable E" by (rule Int_stable_pair_algebra) + show "Int_stable E" by (rule Int_stable_pair_measure_generator) from M1.sets_into_space have "\A. A \ sets M1 \ {x \ space M1. x \ A} = A" by auto then show "sets E \ sets ?D" - by (auto simp: pair_algebra_def sets_sigma if_distrib + by (auto simp: pair_measure_generator_def sets_sigma if_distrib intro: sigma_sets.Basic intro!: M1.measurable_If) - qed auto + qed (auto simp: pair_measure_def) with `Q \ sets P` have "Q \ sets ?D" by simp then show "?s Q \ borel_measurable M1" by simp qed subsection {* Binary products of $\sigma$-finite measure spaces *} -locale pair_sigma_finite = M1: sigma_finite_measure M1 \1 + M2: sigma_finite_measure M2 \2 - for M1 \1 M2 \2 +locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" sublocale pair_sigma_finite \ pair_sigma_algebra M1 M2 by default +lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ ((A = {} \ B = {}) \ (C = {} \ D = {}))" + by auto + lemma (in pair_sigma_finite) measure_cut_measurable_fst: - assumes "Q \ sets P" shows "(\x. \2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") + assumes "Q \ sets P" shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") proof - have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+ - have M1: "sigma_finite_measure M1 \1" by default - + have M1: "sigma_finite_measure M1" by default from M2.disjoint_sigma_finite guess F .. note F = this + then have "\i. F i \ sets M2" by auto let "?C x i" = "F i \ Pair x -` Q" { fix i let ?R = "M2.restricted_space (F i)" have [simp]: "space M1 \ F i \ space M1 \ space M2 = space M1 \ F i" using F M2.sets_into_space by auto - have "(\x. \2 (Pair x -` (space M1 \ F i \ Q))) \ borel_measurable M1" + let ?R2 = "M2.restricted_space (F i)" + have "(\x. measure ?R2 (Pair x -` (space M1 \ space ?R2 \ Q))) \ borel_measurable M1" proof (intro finite_measure_cut_measurable[OF M1]) - show "finite_measure (M2.restricted_space (F i)) \2" + show "finite_measure ?R2" using F by (intro M2.restricted_to_finite_measure) auto - have "space M1 \ F i \ sets P" - using M1.top F by blast - from sigma_sets_Int[symmetric, - OF this[unfolded pair_sigma_algebra_def sets_sigma]] - show "(space M1 \ F i) \ Q \ sets (sigma (pair_algebra M1 ?R))" - using `Q \ sets P` - using pair_algebra_Int_snd[OF M1.space_closed, of "F i" M2] - by (auto simp: pair_algebra_def sets_sigma) + have "(space M1 \ space ?R2) \ Q \ (op \ (space M1 \ F i)) ` sets P" + using `Q \ sets P` by (auto simp: image_iff) + also have "\ = sigma_sets (space M1 \ F i) ((op \ (space M1 \ F i)) ` sets E)" + unfolding pair_measure_def pair_measure_generator_def sigma_def + using `F i \ sets M2` M2.sets_into_space + by (auto intro!: sigma_sets_Int sigma_sets.Basic) + also have "\ \ sets (M1 \\<^isub>M ?R2)" + using M1.sets_into_space + apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def + intro!: sigma_sets_subseteq) + apply (rule_tac x="a" in exI) + apply (rule_tac x="b \ F i" in exI) + by auto + finally show "(space M1 \ space ?R2) \ Q \ sets (M1 \\<^isub>M ?R2)" . qed moreover have "\x. Pair x -` (space M1 \ F i \ Q) = ?C x i" - using `Q \ sets P` sets_into_space by (auto simp: space_pair_algebra) - ultimately have "(\x. \2 (?C x i)) \ borel_measurable M1" + using `Q \ sets P` sets_into_space by (auto simp: space_pair_measure) + ultimately have "(\x. measure M2 (?C x i)) \ borel_measurable M1" by simp } moreover { fix x - have "(\\<^isub>\i. \2 (?C x i)) = \2 (\i. ?C x i)" + have "(\\<^isub>\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" proof (intro M2.measure_countably_additive) show "range (?C x) \ sets M2" using F `Q \ sets P` by (auto intro!: M2.Int measurable_cut_fst) @@ -494,27 +596,24 @@ qed also have "(\i. ?C x i) = Pair x -` Q" using F sets_into_space `Q \ sets P` - by (auto simp: space_pair_algebra) - finally have "\2 (Pair x -` Q) = (\\<^isub>\i. \2 (?C x i))" + by (auto simp: space_pair_measure) + finally have "measure M2 (Pair x -` Q) = (\\<^isub>\i. measure M2 (?C x i))" by simp } ultimately show ?thesis by (auto intro!: M1.borel_measurable_psuminf) qed lemma (in pair_sigma_finite) measure_cut_measurable_snd: - assumes "Q \ sets P" shows "(\y. \1 ((\x. (x, y)) -` Q)) \ borel_measurable M2" + assumes "Q \ sets P" shows "(\y. M1.\ ((\x. (x, y)) -` Q)) \ borel_measurable M2" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default - have [simp]: "\y. (Pair y -` (\(x, y). (y, x)) ` Q) = (\x. (x, y)) -` Q" - by auto + interpret Q: pair_sigma_finite M2 M1 by default note sets_pair_sigma_algebra_swap[OF assms] from Q.measure_cut_measurable_fst[OF this] - show ?thesis by simp + show ?thesis by (simp add: vimage_compose[symmetric] comp_def) qed lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable: - assumes "f \ measurable P M" - shows "(\(x,y). f (y, x)) \ measurable (sigma (pair_algebra M2 M1)) M" + assumes "f \ measurable P M" shows "(\(x,y). f (y, x)) \ measurable (M2 \\<^isub>M M1) M" proof - interpret Q: pair_sigma_algebra M2 M1 by default have *: "(\(x,y). f (y, x)) = f \ (\(x,y). (y, x))" by (simp add: fun_eq_iff) @@ -523,19 +622,15 @@ unfolding * by (rule measurable_comp) qed -definition (in pair_sigma_finite) - "pair_measure A = M1.positive_integral (\x. - M2.positive_integral (\y. indicator A (x, y)))" - lemma (in pair_sigma_finite) pair_measure_alt: assumes "A \ sets P" - shows "pair_measure A = M1.positive_integral (\x. \2 (Pair x -` A))" - unfolding pair_measure_def + shows "measure (M1 \\<^isub>M M2) A = (\\<^isup>+ x. measure M2 (Pair x -` A) \M1)" + apply (simp add: pair_measure_def pair_measure_generator_def) proof (rule M1.positive_integral_cong) fix x assume "x \ space M1" have *: "\y. indicator A (x, y) = (indicator (Pair x -` A) y :: pextreal)" unfolding indicator_def by auto - show "M2.positive_integral (\y. indicator A (x, y)) = \2 (Pair x -` A)" + show "(\\<^isup>+ y. indicator A (x, y) \M2) = measure M2 (Pair x -` A)" unfolding * apply (subst M2.positive_integral_indicator) apply (rule measurable_cut_fst[OF assms]) @@ -544,30 +639,29 @@ lemma (in pair_sigma_finite) pair_measure_times: assumes A: "A \ sets M1" and "B \ sets M2" - shows "pair_measure (A \ B) = \1 A * \2 B" + shows "measure (M1 \\<^isub>M M2) (A \ B) = M1.\ A * measure M2 B" proof - - from assms have "pair_measure (A \ B) = - M1.positive_integral (\x. \2 B * indicator A x)" - by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt) + have "measure (M1 \\<^isub>M M2) (A \ B) = (\\<^isup>+ x. measure M2 B * indicator A x \M1)" + using assms by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt) with assms show ?thesis by (simp add: M1.positive_integral_cmult_indicator ac_simps) qed -lemma (in pair_sigma_finite) sigma_finite_up_in_pair_algebra: +lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: "\F::nat \ ('a \ 'b) set. range F \ sets E \ F \ space E \ - (\i. pair_measure (F i) \ \)" + (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" proof - obtain F1 :: "nat \ 'a set" and F2 :: "nat \ 'b set" where - F1: "range F1 \ sets M1" "F1 \ space M1" "\i. \1 (F1 i) \ \" and - F2: "range F2 \ sets M2" "F2 \ space M2" "\i. \2 (F2 i) \ \" + F1: "range F1 \ sets M1" "F1 \ space M1" "\i. M1.\ (F1 i) \ \" and + F2: "range F2 \ sets M2" "F2 \ space M2" "\i. M2.\ (F2 i) \ \" using M1.sigma_finite_up M2.sigma_finite_up by auto then have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" unfolding isoton_def by auto let ?F = "\i. F1 i \ F2 i" - show ?thesis unfolding isoton_def space_pair_algebra + show ?thesis unfolding isoton_def space_pair_measure proof (intro exI[of _ ?F] conjI allI) show "range ?F \ sets E" using F1 F2 - by (fastsimp intro!: pair_algebraI) + by (fastsimp intro!: pair_measure_generatorI) next have "space M1 \ space M2 \ (\i. ?F i)" proof (intro subsetI) @@ -581,8 +675,8 @@ by (intro SigmaI) (auto simp add: min_max.sup_commute) then show "x \ (\i. ?F i)" by auto qed - then show "(\i. ?F i) = space M1 \ space M2" - using space by (auto simp: space) + then show "(\i. ?F i) = space E" + using space by (auto simp: space pair_measure_generator_def) next fix i show "F1 i \ F2 i \ F1 (Suc i) \ F2 (Suc i)" using `F1 \ space M1` `F2 \ space M2` unfolding isoton_def @@ -590,103 +684,163 @@ next fix i from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto - with F1 F2 show "pair_measure (F1 i \ F2 i) \ \" + with F1 F2 show "measure P (F1 i \ F2 i) \ \" by (simp add: pair_measure_times) qed qed -sublocale pair_sigma_finite \ sigma_finite_measure P pair_measure +sublocale pair_sigma_finite \ sigma_finite_measure P proof - show "pair_measure {} = 0" - unfolding pair_measure_def by auto + show "measure P {} = 0" + unfolding pair_measure_def pair_measure_generator_def sigma_def by auto - show "countably_additive P pair_measure" + show "countably_additive P (measure P)" unfolding countably_additive_def proof (intro allI impI) fix F :: "nat \ ('a \ 'b) set" assume F: "range F \ sets P" "disjoint_family F" from F have *: "\i. F i \ sets P" "(\i. F i) \ sets P" by auto - moreover from F have "\i. (\x. \2 (Pair x -` F i)) \ borel_measurable M1" + moreover from F have "\i. (\x. measure M2 (Pair x -` F i)) \ borel_measurable M1" by (intro measure_cut_measurable_fst) auto moreover have "\x. disjoint_family (\i. Pair x -` F i)" by (intro disjoint_family_on_bisimulation[OF F(2)]) auto moreover have "\x. x \ space M1 \ range (\i. Pair x -` F i) \ sets M2" using F by (auto intro!: measurable_cut_fst) - ultimately show "(\\<^isub>\n. pair_measure (F n)) = pair_measure (\i. F i)" + ultimately show "(\\<^isub>\n. measure P (F n)) = measure P (\i. F i)" by (simp add: pair_measure_alt vimage_UN M1.positive_integral_psuminf[symmetric] M2.measure_countably_additive cong: M1.positive_integral_cong) qed - from sigma_finite_up_in_pair_algebra guess F :: "nat \ ('a \ 'c) set" .. note F = this - show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. pair_measure (F i) \ \)" + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this + show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. measure P (F i) \ \)" proof (rule exI[of _ F], intro conjI) - show "range F \ sets P" using F by auto + show "range F \ sets P" using F by (auto simp: pair_measure_def) show "(\i. F i) = space P" - using F by (auto simp: space_pair_algebra isoton_def) - show "\i. pair_measure (F i) \ \" using F by auto + using F by (auto simp: pair_measure_def pair_measure_generator_def isoton_def) + show "\i. measure P (F i) \ \" using F by auto qed qed lemma (in pair_sigma_algebra) sets_swap: assumes "A \ sets P" - shows "(\(x, y). (y, x)) -` A \ space (sigma (pair_algebra M2 M1)) \ sets (sigma (pair_algebra M2 M1))" + shows "(\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1) \ sets (M2 \\<^isub>M M1)" (is "_ -` A \ space ?Q \ sets ?Q") proof - - have *: "(\(x, y). (y, x)) -` A \ space ?Q = (\(x, y). (y, x)) ` A" - using `A \ sets P` sets_into_space by (auto simp: space_pair_algebra) + have *: "(\(x, y). (y, x)) -` A \ space ?Q = (\(x, y). (y, x)) -` A" + using `A \ sets P` sets_into_space by (auto simp: space_pair_measure) show ?thesis unfolding * using assms by (rule sets_pair_sigma_algebra_swap) qed lemma (in pair_sigma_finite) pair_measure_alt2: assumes "A \ sets P" - shows "pair_measure A = M2.positive_integral (\y. \1 ((\x. (x, y)) -` A))" + shows "\ A = (\\<^isup>+y. M1.\ ((\x. (x, y)) -` A) \M2)" (is "_ = ?\ A") proof - - from sigma_finite_up_in_pair_algebra guess F :: "nat \ ('a \ 'c) set" .. note F = this + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this + have [simp]: "\m. \ space = space E, sets = sets (sigma E), measure = m \ = P\ measure := m \" + unfolding pair_measure_def by simp show ?thesis - proof (rule measure_unique_Int_stable[where \="?\", OF Int_stable_pair_algebra], - simp_all add: pair_sigma_algebra_def[symmetric]) - show "range F \ sets E" "F \ space E" "\i. pair_measure (F i) \ \" - using F by auto - show "measure_space P pair_measure" by default - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default - have P: "sigma_algebra P" by default - show "measure_space P ?\" + proof (rule measure_unique_Int_stable[where \="?\", OF Int_stable_pair_measure_generator], simp_all) + show "range F \ sets E" "F \ space E" "\i. \ (F i) \ \" "A \ sets (sigma E)" + using F `A \ sets P` by (auto simp: pair_measure_def) + show "measure_space P" by default + interpret Q: pair_sigma_finite M2 M1 by default + have P: "sigma_algebra (P\ measure := ?\\)" + by (intro sigma_algebra_cong) auto + show "measure_space (P\ measure := ?\\)" apply (rule Q.measure_space_vimage[OF P]) + apply (simp_all) apply (rule Q.pair_sigma_algebra_swap_measurable) proof - fix A assume "A \ sets P" from sets_swap[OF this] - show "M2.positive_integral (\y. \1 ((\x. (x, y)) -` A)) = - Q.pair_measure ((\(x, y). (y, x)) -` A \ space Q.P)" + show "(\\<^isup>+ y. M1.\ ((\x. (x, y)) -` A) \M2) = Q.\ ((\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1))" using sets_into_space[OF `A \ sets P`] - by (auto simp add: Q.pair_measure_alt space_pair_algebra - intro!: M2.positive_integral_cong arg_cong[where f=\1]) + by (auto simp add: Q.pair_measure_alt space_pair_measure + intro!: M2.positive_integral_cong arg_cong[where f="M1.\"]) qed fix X assume "X \ sets E" then obtain A B where X: "X = A \ B" and AB: "A \ sets M1" "B \ sets M2" - unfolding pair_algebra_def by auto - show "pair_measure X = ?\ X" + unfolding pair_measure_def pair_measure_generator_def by auto + show "\ X = ?\ X" proof - - from AB have "?\ (A \ B) = - M2.positive_integral (\y. \1 A * indicator B y)" + from AB have "?\ (A \ B) = (\\<^isup>+y. M1.\ A * indicator B y \M2)" by (auto intro!: M2.positive_integral_cong) with AB show ?thesis unfolding pair_measure_times[OF AB] X by (simp add: M2.positive_integral_cmult_indicator ac_simps) qed - qed fact + qed +qed + + +lemma pair_sigma_algebra_sigma: + assumes 1: "S1 \ (space E1)" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" + assumes 2: "S2 \ (space E2)" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" + shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))" + (is "sets ?S = sets ?E") +proof - + interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma) + interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma) + have P: "sets (pair_measure_generator E1 E2) \ Pow (space E1 \ space E2)" + using E1 E2 by (auto simp add: pair_measure_generator_def) + interpret E: sigma_algebra ?E unfolding pair_measure_generator_def + using E1 E2 by (intro sigma_algebra_sigma) auto + { fix A assume "A \ sets E1" + then have "fst -` A \ space ?E = A \ (\i. S2 i)" + using E1 2 unfolding isoton_def pair_measure_generator_def by auto + also have "\ = (\i. A \ S2 i)" by auto + also have "\ \ sets ?E" unfolding pair_measure_generator_def sets_sigma + using 2 `A \ sets E1` + by (intro sigma_sets.Union) + (auto simp: image_subset_iff intro!: sigma_sets.Basic) + finally have "fst -` A \ space ?E \ sets ?E" . } + moreover + { fix B assume "B \ sets E2" + then have "snd -` B \ space ?E = (\i. S1 i) \ B" + using E2 1 unfolding isoton_def pair_measure_generator_def by auto + also have "\ = (\i. S1 i \ B)" by auto + also have "\ \ sets ?E" + using 1 `B \ sets E2` unfolding pair_measure_generator_def sets_sigma + by (intro sigma_sets.Union) + (auto simp: image_subset_iff intro!: sigma_sets.Basic) + finally have "snd -` B \ space ?E \ sets ?E" . } + ultimately have proj: + "fst \ measurable ?E (sigma E1) \ snd \ measurable ?E (sigma E2)" + using E1 E2 by (subst (1 2) E.measurable_iff_sigma) + (auto simp: pair_measure_generator_def sets_sigma) + { fix A B assume A: "A \ sets (sigma E1)" and B: "B \ sets (sigma E2)" + with proj have "fst -` A \ space ?E \ sets ?E" "snd -` B \ space ?E \ sets ?E" + unfolding measurable_def by simp_all + moreover have "A \ B = (fst -` A \ space ?E) \ (snd -` B \ space ?E)" + using A B M1.sets_into_space M2.sets_into_space + by (auto simp: pair_measure_generator_def) + ultimately have "A \ B \ sets ?E" by auto } + then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \ sets ?E" + by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma) + then have subset: "sets ?S \ sets ?E" + by (simp add: sets_sigma pair_measure_generator_def) + show "sets ?S = sets ?E" + proof (intro set_eqI iffI) + fix A assume "A \ sets ?E" then show "A \ sets ?S" + unfolding sets_sigma + proof induct + case (Basic A) then show ?case + by (auto simp: pair_measure_generator_def sets_sigma intro: sigma_sets.Basic) + qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def) + next + fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto + qed qed section "Fubinis theorem" lemma (in pair_sigma_finite) simple_function_cut: - assumes f: "simple_function f" - shows "(\x. M2.positive_integral (\ y. f (x, y))) \ borel_measurable M1" - and "M1.positive_integral (\x. M2.positive_integral (\y. f (x, y))) - = positive_integral f" + assumes f: "simple_function P f" + shows "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" proof - have f_borel: "f \ borel_measurable P" using f by (rule borel_measurable_simple_function) @@ -696,42 +850,40 @@ have [simp]: "\z y. indicator (?F z) (x, y) = indicator (?F' x z) y" by (auto simp: indicator_def) have "\y. y \ space M2 \ (x, y) \ space P" using `x \ space M1` - by (simp add: space_pair_algebra) + by (simp add: space_pair_measure) moreover have "\x z. ?F' x z \ sets M2" using f_borel by (intro borel_measurable_vimage measurable_cut_fst) - ultimately have "M2.simple_function (\ y. f (x, y))" + ultimately have "simple_function M2 (\ y. f (x, y))" apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _]) apply (rule simple_function_indicator_representation[OF f]) using `x \ space M1` by (auto simp del: space_sigma) } note M2_sf = this { fix x assume x: "x \ space M1" - then have "M2.positive_integral (\ y. f (x, y)) = - (\z\f ` space P. z * \2 (?F' x z))" + then have "(\\<^isup>+y. f (x, y) \M2) = (\z\f ` space P. z * M2.\ (?F' x z))" unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x]] - unfolding M2.simple_integral_def + unfolding simple_integral_def proof (safe intro!: setsum_mono_zero_cong_left) from f show "finite (f ` space P)" by (rule simple_functionD) next fix y assume "y \ space M2" then show "f (x, y) \ f ` space P" - using `x \ space M1` by (auto simp: space_pair_algebra) + using `x \ space M1` by (auto simp: space_pair_measure) next fix x' y assume "(x', y) \ space P" "f (x', y) \ (\y. f (x, y)) ` space M2" then have *: "?F' x (f (x', y)) = {}" - by (force simp: space_pair_algebra) - show "f (x', y) * \2 (?F' x (f (x', y))) = 0" + by (force simp: space_pair_measure) + show "f (x', y) * M2.\ (?F' x (f (x', y))) = 0" unfolding * by simp qed (simp add: vimage_compose[symmetric] comp_def - space_pair_algebra) } + space_pair_measure) } note eq = this moreover have "\z. ?F z \ sets P" by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma) - moreover then have "\z. (\x. \2 (?F' x z)) \ borel_measurable M1" + moreover then have "\z. (\x. M2.\ (?F' x z)) \ borel_measurable M1" by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int) ultimately - show "(\ x. M2.positive_integral (\ y. f (x, y))) \ borel_measurable M1" - and "M1.positive_integral (\x. M2.positive_integral (\y. f (x, y))) - = positive_integral f" + show "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" by (auto simp del: vimage_Int cong: measurable_cong intro!: M1.borel_measurable_pextreal_setsum simp add: M1.positive_integral_setsum simple_integral_def @@ -743,13 +895,12 @@ lemma (in pair_sigma_finite) positive_integral_fst_measurable: assumes f: "f \ borel_measurable P" - shows "(\ x. M2.positive_integral (\ y. f (x, y))) \ borel_measurable M1" + shows "(\x. \\<^isup>+ y. f (x, y) \M2) \ borel_measurable M1" (is "?C f \ borel_measurable M1") - and "M1.positive_integral (\ x. M2.positive_integral (\ y. f (x, y))) = - positive_integral f" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" proof - from borel_measurable_implies_simple_function_sequence[OF f] - obtain F where F: "\i. simple_function (F i)" "F \ f" by auto + obtain F where F: "\i. simple_function P (F i)" "F \ f" by auto then have F_borel: "\i. F i \ borel_measurable P" and F_mono: "\i x. F i x \ F (Suc i) x" and F_SUPR: "\x. (SUP i. F i x) = f x" @@ -770,56 +921,48 @@ note SUPR_C = this ultimately show "?C f \ borel_measurable M1" by (simp cong: measurable_cong) - have "positive_integral (\x. (SUP i. F i x)) = (SUP i. positive_integral (F i))" + have "(\\<^isup>+x. (SUP i. F i x) \P) = (SUP i. integral\<^isup>P P (F i))" using F_borel F_mono by (auto intro!: positive_integral_monotone_convergence_SUP[symmetric]) - also have "(SUP i. positive_integral (F i)) = - (SUP i. M1.positive_integral (\x. M2.positive_integral (\y. F i (x, y))))" + also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" unfolding sf(2) by simp - also have "\ = M1.positive_integral (\x. SUP i. M2.positive_integral (\y. F i (x, y)))" + also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" by (auto intro!: M1.positive_integral_monotone_convergence_SUP[OF _ sf(1)] M2.positive_integral_mono F_mono) - also have "\ = M1.positive_integral (\x. M2.positive_integral (\y. SUP i. F i (x, y)))" + also have "\ = \\<^isup>+ x. (\\<^isup>+ y. (SUP i. F i (x, y)) \M2) \M1" using F_borel F_mono by (auto intro!: M2.positive_integral_monotone_convergence_SUP M1.positive_integral_cong measurable_pair_image_snd) - finally show "M1.positive_integral (\ x. M2.positive_integral (\ y. f (x, y))) = - positive_integral f" + finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" unfolding F_SUPR by simp qed lemma (in pair_sigma_finite) positive_integral_product_swap: assumes f: "f \ borel_measurable P" - shows "measure_space.positive_integral - (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \2 M1 \1) (\x. f (case x of (x,y)\(y,x))) = - positive_integral f" + shows "(\\<^isup>+x. f (case x of (x,y)\(y,x)) \(M2 \\<^isub>M M1)) = integral\<^isup>P P f" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default - have P: "sigma_algebra P" by default + interpret Q: pair_sigma_finite M2 M1 by default + have "sigma_algebra P" by default show ?thesis - unfolding Q.positive_integral_vimage[OF P Q.pair_sigma_algebra_swap_measurable f, symmetric] - proof (rule positive_integral_cong_measure) - fix A - assume A: "A \ sets P" + proof (intro Q.positive_integral_vimage[symmetric] Q.pair_sigma_algebra_swap_measurable) + fix A assume "A \ sets P" from Q.pair_sigma_algebra_swap_measurable[THEN measurable_sets, OF this] this sets_into_space[OF this] - show "Q.pair_measure ((\(x, y). (y, x)) -` A \ space Q.P) = pair_measure A" - by (auto intro!: M1.positive_integral_cong arg_cong[where f=\2] - simp: pair_measure_alt Q.pair_measure_alt2 space_pair_algebra) - qed + show "\ A = Q.\ ((\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1))" + by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\"] + simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure) + qed fact+ qed lemma (in pair_sigma_finite) positive_integral_snd_measurable: assumes f: "f \ borel_measurable P" - shows "M2.positive_integral (\y. M1.positive_integral (\x. f (x, y))) = - positive_integral f" + shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = integral\<^isup>P P f" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default + interpret Q: pair_sigma_finite M2 M1 by default note pair_sigma_algebra_measurable[OF f] from Q.positive_integral_fst_measurable[OF this] - have "M2.positive_integral (\y. M1.positive_integral (\x. f (x, y))) = - Q.positive_integral (\(x, y). f (y, x))" + have "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ (x, y). f (y, x) \Q.P)" by simp - also have "Q.positive_integral (\(x, y). f (y, x)) = positive_integral f" + also have "(\\<^isup>+ (x, y). f (y, x) \Q.P) = integral\<^isup>P P f" unfolding positive_integral_product_swap[OF f, symmetric] by (auto intro!: Q.positive_integral_cong) finally show ?thesis . @@ -827,8 +970,7 @@ lemma (in pair_sigma_finite) Fubini: assumes f: "f \ borel_measurable P" - shows "M2.positive_integral (\y. M1.positive_integral (\x. f (x, y))) = - M1.positive_integral (\x. M2.positive_integral (\y. f (x, y)))" + shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1)" unfolding positive_integral_snd_measurable[OF assms] unfolding positive_integral_fst_measurable[OF assms] .. @@ -836,30 +978,30 @@ assumes "almost_everywhere (\x. Q x)" shows "M1.almost_everywhere (\x. M2.almost_everywhere (\y. Q (x, y)))" proof - - obtain N where N: "N \ sets P" "pair_measure N = 0" "{x\space P. \ Q x} \ N" + obtain N where N: "N \ sets P" "\ N = 0" "{x\space P. \ Q x} \ N" using assms unfolding almost_everywhere_def by auto show ?thesis proof (rule M1.AE_I) from N measure_cut_measurable_fst[OF `N \ sets P`] - show "\1 {x\space M1. \2 (Pair x -` N) \ 0} = 0" + show "M1.\ {x\space M1. M2.\ (Pair x -` N) \ 0} = 0" by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def) - show "{x \ space M1. \2 (Pair x -` N) \ 0} \ sets M1" + show "{x \ space M1. M2.\ (Pair x -` N) \ 0} \ sets M1" by (intro M1.borel_measurable_pextreal_neq_const measure_cut_measurable_fst N) - { fix x assume "x \ space M1" "\2 (Pair x -` N) = 0" + { fix x assume "x \ space M1" "M2.\ (Pair x -` N) = 0" have "M2.almost_everywhere (\y. Q (x, y))" proof (rule M2.AE_I) - show "\2 (Pair x -` N) = 0" by fact + show "M2.\ (Pair x -` N) = 0" by fact show "Pair x -` N \ sets M2" by (intro measurable_cut_fst N) show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" - using N `x \ space M1` unfolding space_sigma space_pair_algebra by auto + using N `x \ space M1` unfolding space_sigma space_pair_measure by auto qed } - then show "{x \ space M1. \ M2.almost_everywhere (\y. Q (x, y))} \ {x \ space M1. \2 (Pair x -` N) \ 0}" + then show "{x \ space M1. \ M2.almost_everywhere (\y. Q (x, y))} \ {x \ space M1. M2.\ (Pair x -` N) \ 0}" by auto qed qed lemma (in pair_sigma_algebra) measurable_product_swap: - "f \ measurable (sigma (pair_algebra M2 M1)) M \ (\(x,y). f (y,x)) \ measurable P M" + "f \ measurable (M2 \\<^isub>M M1) M \ (\(x,y). f (y,x)) \ measurable P M" proof - interpret Q: pair_sigma_algebra M2 M1 by default show ?thesis @@ -868,114 +1010,106 @@ qed lemma (in pair_sigma_finite) integrable_product_swap: - assumes "integrable f" - shows "measure_space.integrable - (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \2 M1 \1) (\(x,y). f (y,x))" + assumes "integrable P f" + shows "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x))" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default + interpret Q: pair_sigma_finite M2 M1 by default have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * - using assms unfolding Q.integrable_def integrable_def + using assms unfolding integrable_def apply (subst (1 2) positive_integral_product_swap) - using `integrable f` unfolding integrable_def + using `integrable P f` unfolding integrable_def by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) qed lemma (in pair_sigma_finite) integrable_product_swap_iff: - "measure_space.integrable - (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \2 M1 \1) (\(x,y). f (y,x)) \ - integrable f" + "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x)) \ integrable P f" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default + interpret Q: pair_sigma_finite M2 M1 by default from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] show ?thesis by auto qed lemma (in pair_sigma_finite) integral_product_swap: - assumes "integrable f" - shows "measure_space.integral - (sigma (pair_algebra M2 M1)) (pair_sigma_finite.pair_measure M2 \2 M1 \1) (\(x,y). f (y,x)) = - integral f" + assumes "integrable P f" + shows "(\(x,y). f (y,x) \(M2 \\<^isub>M M1)) = integral\<^isup>L P f" proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default + interpret Q: pair_sigma_finite M2 M1 by default have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis - unfolding integral_def Q.integral_def * + unfolding lebesgue_integral_def * apply (subst (1 2) positive_integral_product_swap) - using `integrable f` unfolding integrable_def + using `integrable P f` unfolding integrable_def by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) qed lemma (in pair_sigma_finite) integrable_fst_measurable: - assumes f: "integrable f" - shows "M1.almost_everywhere (\x. M2.integrable (\ y. f (x, y)))" (is "?AE") - and "M1.integral (\ x. M2.integral (\ y. f (x, y))) = integral f" (is "?INT") + assumes f: "integrable P f" + shows "M1.almost_everywhere (\x. integrable M2 (\ y. f (x, y)))" (is "?AE") + and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^isup>L P f" (is "?INT") proof - let "?pf x" = "Real (f x)" and "?nf x" = "Real (- f x)" have borel: "?nf \ borel_measurable P""?pf \ borel_measurable P" and - int: "positive_integral ?nf \ \" "positive_integral ?pf \ \" + int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" using assms by auto - have "M1.positive_integral (\x. M2.positive_integral (\y. Real (f (x, y)))) \ \" - "M1.positive_integral (\x. M2.positive_integral (\y. Real (- f (x, y)))) \ \" + have "(\\<^isup>+x. (\\<^isup>+y. Real (f (x, y)) \M2) \M1) \ \" + "(\\<^isup>+x. (\\<^isup>+y. Real (- f (x, y)) \M2) \M1) \ \" using borel[THEN positive_integral_fst_measurable(1)] int unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all with borel[THEN positive_integral_fst_measurable(1)] - have AE: "M1.almost_everywhere (\x. M2.positive_integral (\y. Real (f (x, y))) \ \)" - "M1.almost_everywhere (\x. M2.positive_integral (\y. Real (- f (x, y))) \ \)" + have AE: "M1.almost_everywhere (\x. (\\<^isup>+y. Real (f (x, y)) \M2) \ \)" + "M1.almost_everywhere (\x. (\\<^isup>+y. Real (- f (x, y)) \M2) \ \)" by (auto intro!: M1.positive_integral_omega_AE) then show ?AE apply (rule M1.AE_mp[OF _ M1.AE_mp]) apply (rule M1.AE_cong) - using assms unfolding M2.integrable_def + using assms unfolding integrable_def by (auto intro!: measurable_pair_image_snd) - have "M1.integrable - (\x. real (M2.positive_integral (\xa. Real (f (x, xa)))))" (is "M1.integrable ?f") - proof (unfold M1.integrable_def, intro conjI) + have "integrable M1 (\x. real (\\<^isup>+y. Real (f (x, y)) \M2))" (is "integrable M1 ?f") + proof (intro integrable_def[THEN iffD2] conjI) show "?f \ borel_measurable M1" using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable) - have "M1.positive_integral (\x. Real (?f x)) = - M1.positive_integral (\x. M2.positive_integral (\xa. Real (f (x, xa))))" + have "(\\<^isup>+x. Real (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. Real (f (x, y)) \M2) \M1)" apply (rule M1.positive_integral_cong_AE) apply (rule M1.AE_mp[OF AE(1)]) apply (rule M1.AE_cong) by (auto simp: Real_real) - then show "M1.positive_integral (\x. Real (?f x)) \ \" + then show "(\\<^isup>+x. Real (?f x) \M1) \ \" using positive_integral_fst_measurable[OF borel(2)] int by simp - have "M1.positive_integral (\x. Real (- ?f x)) = M1.positive_integral (\x. 0)" + have "(\\<^isup>+x. Real (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" by (intro M1.positive_integral_cong) simp - then show "M1.positive_integral (\x. Real (- ?f x)) \ \" by simp + then show "(\\<^isup>+x. Real (- ?f x) \M1) \ \" by simp qed - moreover have "M1.integrable - (\x. real (M2.positive_integral (\xa. Real (- f (x, xa)))))" (is "M1.integrable ?f") - proof (unfold M1.integrable_def, intro conjI) + moreover have "integrable M1 (\x. real (\\<^isup>+ y. Real (- f (x, y)) \M2))" + (is "integrable M1 ?f") + proof (intro integrable_def[THEN iffD2] conjI) show "?f \ borel_measurable M1" using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable) - have "M1.positive_integral (\x. Real (?f x)) = - M1.positive_integral (\x. M2.positive_integral (\xa. Real (- f (x, xa))))" + have "(\\<^isup>+x. Real (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. Real (- f (x, y)) \M2) \M1)" apply (rule M1.positive_integral_cong_AE) apply (rule M1.AE_mp[OF AE(2)]) apply (rule M1.AE_cong) by (auto simp: Real_real) - then show "M1.positive_integral (\x. Real (?f x)) \ \" + then show "(\\<^isup>+x. Real (?f x) \M1) \ \" using positive_integral_fst_measurable[OF borel(1)] int by simp - have "M1.positive_integral (\x. Real (- ?f x)) = M1.positive_integral (\x. 0)" + have "(\\<^isup>+x. Real (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" by (intro M1.positive_integral_cong) simp - then show "M1.positive_integral (\x. Real (- ?f x)) \ \" by simp + then show "(\\<^isup>+x. Real (- ?f x) \M1) \ \" by simp qed ultimately show ?INT - unfolding M2.integral_def integral_def + unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2] borel[THEN positive_integral_fst_measurable(2), symmetric] by (simp add: M1.integral_real[OF AE(1)] M1.integral_real[OF AE(2)]) qed lemma (in pair_sigma_finite) integrable_snd_measurable: - assumes f: "integrable f" - shows "M2.almost_everywhere (\y. M1.integrable (\x. f (x, y)))" (is "?AE") - and "M2.integral (\y. M1.integral (\x. f (x, y))) = integral f" (is "?INT") + assumes f: "integrable P f" + shows "M2.almost_everywhere (\y. integrable M1 (\x. f (x, y)))" (is "?AE") + and "(\y. (\x. f (x, y) \M1) \M2) = integral\<^isup>L P f" (is "?INT") proof - - interpret Q: pair_sigma_finite M2 \2 M1 \1 by default - have Q_int: "Q.integrable (\(x, y). f (y, x))" + interpret Q: pair_sigma_finite M2 M1 by default + have Q_int: "integrable Q.P (\(x, y). f (y, x))" using f unfolding integrable_product_swap_iff . show ?INT using Q.integrable_fst_measurable(2)[OF Q_int] @@ -986,9 +1120,8 @@ qed lemma (in pair_sigma_finite) Fubini_integral: - assumes f: "integrable f" - shows "M2.integral (\y. M1.integral (\x. f (x, y))) = - M1.integral (\x. M2.integral (\y. f (x, y)))" + assumes f: "integrable P f" + shows "(\y. (\x. f (x, y) \M1) \M2) = (\x. (\y. f (x, y) \M2) \M1)" unfolding integrable_snd_measurable[OF assms] unfolding integrable_fst_measurable[OF assms] .. @@ -997,291 +1130,140 @@ section "Products" locale product_sigma_algebra = - fixes M :: "'i \ 'a algebra" + fixes M :: "'i \ ('a, 'b) measure_space_scheme" assumes sigma_algebras: "\i. sigma_algebra (M i)" -locale finite_product_sigma_algebra = product_sigma_algebra M for M :: "'i \ 'a algebra" + +locale finite_product_sigma_algebra = product_sigma_algebra M + for M :: "'i \ ('a, 'b) measure_space_scheme" + fixes I :: "'i set" assumes finite_index: "finite I" +definition + "product_algebra_generator I M = \ space = (\\<^isub>E i \ I. space (M i)), + sets = Pi\<^isub>E I ` (\ i \ I. sets (M i)), + measure = \A. (\i\I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \" + +definition product_algebra_def: + "Pi\<^isub>M I M = sigma (product_algebra_generator I M) + \ measure := (SOME \. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ + (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. measure (M i) (A i))))\" + syntax - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10) + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10) syntax (xsymbols) - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) syntax (HTML output) - "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\\<^isub>E _\_./ _)" 10) + "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => + ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) translations - "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" + "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" -definition - "product_algebra M I = \ space = (\\<^isub>E i\I. space (M i)), sets = Pi\<^isub>E I ` (\ i \ I. sets (M i)) \" - -abbreviation (in finite_product_sigma_algebra) "G \ product_algebra M I" -abbreviation (in finite_product_sigma_algebra) "P \ sigma G" +abbreviation (in finite_product_sigma_algebra) "G \ product_algebra_generator I M" +abbreviation (in finite_product_sigma_algebra) "P \ Pi\<^isub>M I M" sublocale product_sigma_algebra \ M: sigma_algebra "M i" for i by (rule sigma_algebras) -lemma (in finite_product_sigma_algebra) product_algebra_into_space: - "sets G \ Pow (space G)" - using M.sets_into_space unfolding product_algebra_def +lemma sigma_into_space: + assumes "sets M \ Pow (space M)" + shows "sets (sigma M) \ Pow (space M)" + using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto + +lemma (in product_sigma_algebra) product_algebra_generator_into_space: + "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" + using M.sets_into_space unfolding product_algebra_generator_def by auto blast +lemma (in product_sigma_algebra) product_algebra_into_space: + "sets (Pi\<^isub>M I M) \ Pow (space (Pi\<^isub>M I M))" + using product_algebra_generator_into_space + by (auto intro!: sigma_into_space simp add: product_algebra_def) + +lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)" + using product_algebra_generator_into_space unfolding product_algebra_def + by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all + sublocale finite_product_sigma_algebra \ sigma_algebra P - using product_algebra_into_space by (rule sigma_algebra_sigma) + using sigma_algebra_product_algebra . lemma product_algebraE: - assumes "A \ sets (product_algebra M I)" + assumes "A \ sets (product_algebra_generator I M)" obtains E where "A = Pi\<^isub>E I E" "E \ (\ i\I. sets (M i))" - using assms unfolding product_algebra_def by auto + using assms unfolding product_algebra_generator_def by auto -lemma product_algebraI[intro]: +lemma product_algebra_generatorI[intro]: assumes "E \ (\ i\I. sets (M i))" - shows "Pi\<^isub>E I E \ sets (product_algebra M I)" - using assms unfolding product_algebra_def by auto + shows "Pi\<^isub>E I E \ sets (product_algebra_generator I M)" + using assms unfolding product_algebra_generator_def by auto + +lemma space_product_algebra_generator[simp]: + "space (product_algebra_generator I M) = Pi\<^isub>E I (\i. space (M i))" + unfolding product_algebra_generator_def by simp lemma space_product_algebra[simp]: - "space (product_algebra M I) = Pi\<^isub>E I (\i. space (M i))" - unfolding product_algebra_def by simp + "space (Pi\<^isub>M I M) = (\\<^isub>E i\I. space (M i))" + unfolding product_algebra_def product_algebra_generator_def by simp -lemma product_algebra_sets_into_space: +lemma sets_product_algebra: + "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))" + unfolding product_algebra_def sigma_def by simp + +lemma product_algebra_generator_sets_into_space: assumes "\i. i\I \ sets (M i) \ Pow (space (M i))" - shows "sets (product_algebra M I) \ Pow (space (product_algebra M I))" - using assms by (auto simp: product_algebra_def) blast - -lemma (in finite_product_sigma_algebra) P_empty: - "I = {} \ P = \ space = {\k. undefined}, sets = { {}, {\k. undefined} }\" - unfolding product_algebra_def by (simp add: sigma_def image_constant) + shows "sets (product_algebra_generator I M) \ Pow (space (product_algebra_generator I M))" + using assms by (auto simp: product_algebra_generator_def) blast lemma (in finite_product_sigma_algebra) in_P[simp, intro]: "\ \i. i \ I \ A i \ sets (M i) \ \ Pi\<^isub>E I A \ sets P" - by (auto simp: product_algebra_def sets_sigma intro!: sigma_sets.Basic) - -lemma (in product_sigma_algebra) bij_inv_restrict_merge: - assumes [simp]: "I \ J = {}" - shows "bij_inv - (space (sigma (product_algebra M (I \ J)))) - (space (sigma (pair_algebra (product_algebra M I) (product_algebra M J)))) - (\x. (restrict x I, restrict x J)) (\(x, y). merge I x J y)" - by (rule bij_invI) (auto simp: space_pair_algebra extensional_restrict) - -lemma (in product_sigma_algebra) bij_inv_singleton: - "bij_inv (space (sigma (product_algebra M {i}))) (space (M i)) - (\x. x i) (\x. (\j\{i}. x))" - by (rule bij_invI) (auto simp: restrict_def extensional_def fun_eq_iff) - -lemma (in product_sigma_algebra) bij_inv_restrict_insert: - assumes [simp]: "i \ I" - shows "bij_inv - (space (sigma (product_algebra M (insert i I)))) - (space (sigma (pair_algebra (product_algebra M I) (M i)))) - (\x. (restrict x I, x i)) (\(x, y). x(i := y))" - by (rule bij_invI) (auto simp: space_pair_algebra extensional_restrict) - -lemma (in product_sigma_algebra) measurable_restrict_on_generating: - assumes [simp]: "I \ J = {}" - shows "(\x. (restrict x I, restrict x J)) \ measurable - (product_algebra M (I \ J)) - (pair_algebra (product_algebra M I) (product_algebra M J))" - (is "?R \ measurable ?IJ ?P") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "?R \ space ?IJ \ space ?P" by (auto simp: space_pair_algebra) - { fix F E assume "E \ (\ i\I. sets (M i))" "F \ (\ i\J. sets (M i))" - then have "Pi (I \ J) (merge I E J F) \ (\\<^isub>E i\I \ J. space (M i)) = - Pi\<^isub>E (I \ J) (merge I E J F)" - using M.sets_into_space by auto blast+ } - note this[simp] - show "\A. A \ sets ?P \ ?R -` A \ space ?IJ \ sets ?IJ" - by (force elim!: pair_algebraE product_algebraE - simp del: vimage_Int simp: restrict_vimage merge_vimage space_pair_algebra) - qed - -lemma (in product_sigma_algebra) measurable_merge_on_generating: - assumes [simp]: "I \ J = {}" - shows "(\(x, y). merge I x J y) \ measurable - (pair_algebra (product_algebra M I) (product_algebra M J)) - (product_algebra M (I \ J))" - (is "?M \ measurable ?P ?IJ") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "?M \ space ?P \ space ?IJ" by (auto simp: space_pair_algebra) - { fix E assume "E \ (\ i\I. sets (M i))" "E \ (\ i\J. sets (M i))" - then have "Pi I E \ Pi J E \ (\\<^isub>E i\I. space (M i)) \ (\\<^isub>E i\J. space (M i)) = - Pi\<^isub>E I E \ Pi\<^isub>E J E" - using M.sets_into_space by auto blast+ } - note this[simp] - show "\A. A \ sets ?IJ \ ?M -` A \ space ?P \ sets ?P" - by (force elim!: pair_algebraE product_algebraE - simp del: vimage_Int simp: restrict_vimage merge_vimage space_pair_algebra) - qed - -lemma (in product_sigma_algebra) measurable_singleton_on_generator: - "(\x. \j\{i}. x) \ measurable (M i) (product_algebra M {i})" - (is "?f \ measurable _ ?P") -proof (unfold measurable_def, intro CollectI conjI) - show "?f \ space (M i) \ space ?P" by auto - have "\E. E i \ sets (M i) \ ?f -` Pi\<^isub>E {i} E \ space (M i) = E i" - using M.sets_into_space by auto - then show "\A \ sets ?P. ?f -` A \ space (M i) \ sets (M i)" - by (auto elim!: product_algebraE) -qed - -lemma (in product_sigma_algebra) measurable_component_on_generator: - assumes "i \ I" shows "(\x. x i) \ measurable (product_algebra M I) (M i)" - (is "?f \ measurable ?P _") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "?f \ space ?P \ space (M i)" using `i \ I` by auto - fix A assume "A \ sets (M i)" - moreover then have "(\x. x i) -` A \ (\\<^isub>E i\I. space (M i)) = - (\\<^isub>E j\I. if i = j then A else space (M j))" - using M.sets_into_space `i \ I` - by (fastsimp dest: Pi_mem split: split_if_asm) - ultimately show "?f -` A \ space ?P \ sets ?P" - by (auto intro!: product_algebraI) -qed - -lemma (in product_sigma_algebra) measurable_restrict_singleton_on_generating: - assumes [simp]: "i \ I" - shows "(\x. (restrict x I, x i)) \ measurable - (product_algebra M (insert i I)) - (pair_algebra (product_algebra M I) (M i))" - (is "?R \ measurable ?I ?P") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "?R \ space ?I \ space ?P" by (auto simp: space_pair_algebra) - { fix E F assume "E \ (\ i\I. sets (M i))" "F \ sets (M i)" - then have "(\x. (restrict x I, x i)) -` (Pi\<^isub>E I E \ F) \ (\\<^isub>E i\insert i I. space (M i)) = - Pi\<^isub>E (insert i I) (E(i := F))" - using M.sets_into_space using `i\I` by (auto simp: restrict_Pi_cancel) blast+ } - note this[simp] - show "\A. A \ sets ?P \ ?R -` A \ space ?I \ sets ?I" - by (force elim!: pair_algebraE product_algebraE - simp del: vimage_Int simp: space_pair_algebra) -qed - -lemma (in product_sigma_algebra) measurable_merge_singleton_on_generating: - assumes [simp]: "i \ I" - shows "(\(x, y). x(i := y)) \ measurable - (pair_algebra (product_algebra M I) (M i)) - (product_algebra M (insert i I))" - (is "?M \ measurable ?P ?IJ") -proof (unfold measurable_def, intro CollectI conjI ballI) - show "?M \ space ?P \ space ?IJ" by (auto simp: space_pair_algebra) - { fix E assume "E \ (\ i\I. sets (M i))" "E i \ sets (M i)" - then have "(\(x, y). x(i := y)) -` Pi\<^isub>E (insert i I) E \ (\\<^isub>E i\I. space (M i)) \ space (M i) = - Pi\<^isub>E I E \ E i" - using M.sets_into_space by auto blast+ } - note this[simp] - show "\A. A \ sets ?IJ \ ?M -` A \ space ?P \ sets ?P" - by (force elim!: pair_algebraE product_algebraE - simp del: vimage_Int simp: space_pair_algebra) -qed + by (auto simp: sets_product_algebra) section "Generating set generates also product algebra" -lemma pair_sigma_algebra_sigma: - assumes 1: "S1 \ (space E1)" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" - assumes 2: "S2 \ (space E2)" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" - shows "sigma (pair_algebra (sigma E1) (sigma E2)) = sigma (pair_algebra E1 E2)" - (is "?S = ?E") -proof - - interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma) - interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma) - have P: "sets (pair_algebra E1 E2) \ Pow (space E1 \ space E2)" - using E1 E2 by (auto simp add: pair_algebra_def) - interpret E: sigma_algebra ?E unfolding pair_algebra_def - using E1 E2 by (intro sigma_algebra_sigma) auto - { fix A assume "A \ sets E1" - then have "fst -` A \ space ?E = A \ (\i. S2 i)" - using E1 2 unfolding isoton_def pair_algebra_def by auto - also have "\ = (\i. A \ S2 i)" by auto - also have "\ \ sets ?E" unfolding pair_algebra_def sets_sigma - using 2 `A \ sets E1` - by (intro sigma_sets.Union) - (auto simp: image_subset_iff intro!: sigma_sets.Basic) - finally have "fst -` A \ space ?E \ sets ?E" . } - moreover - { fix B assume "B \ sets E2" - then have "snd -` B \ space ?E = (\i. S1 i) \ B" - using E2 1 unfolding isoton_def pair_algebra_def by auto - also have "\ = (\i. S1 i \ B)" by auto - also have "\ \ sets ?E" - using 1 `B \ sets E2` unfolding pair_algebra_def sets_sigma - by (intro sigma_sets.Union) - (auto simp: image_subset_iff intro!: sigma_sets.Basic) - finally have "snd -` B \ space ?E \ sets ?E" . } - ultimately have proj: - "fst \ measurable ?E (sigma E1) \ snd \ measurable ?E (sigma E2)" - using E1 E2 by (subst (1 2) E.measurable_iff_sigma) - (auto simp: pair_algebra_def sets_sigma) - { fix A B assume A: "A \ sets (sigma E1)" and B: "B \ sets (sigma E2)" - with proj have "fst -` A \ space ?E \ sets ?E" "snd -` B \ space ?E \ sets ?E" - unfolding measurable_def by simp_all - moreover have "A \ B = (fst -` A \ space ?E) \ (snd -` B \ space ?E)" - using A B M1.sets_into_space M2.sets_into_space - by (auto simp: pair_algebra_def) - ultimately have "A \ B \ sets ?E" by auto } - then have "sigma_sets (space ?E) (sets (pair_algebra (sigma E1) (sigma E2))) \ sets ?E" - by (intro E.sigma_sets_subset) (auto simp add: pair_algebra_def sets_sigma) - then have subset: "sets ?S \ sets ?E" - by (simp add: sets_sigma pair_algebra_def) - have "sets ?S = sets ?E" - proof (intro set_eqI iffI) - fix A assume "A \ sets ?E" then show "A \ sets ?S" - unfolding sets_sigma - proof induct - case (Basic A) then show ?case - by (auto simp: pair_algebra_def sets_sigma intro: sigma_sets.Basic) - qed (auto intro: sigma_sets.intros simp: pair_algebra_def) - next - fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto - qed - then show ?thesis - by (simp add: pair_algebra_def sigma_def) -qed - lemma sigma_product_algebra_sigma_eq: assumes "finite I" assumes isotone: "\i. i \ I \ (S i) \ (space (E i))" assumes sets_into: "\i. i \ I \ range (S i) \ sets (E i)" and E: "\i. sets (E i) \ Pow (space (E i))" - shows "sigma (product_algebra (\i. sigma (E i)) I) = sigma (product_algebra E I)" - (is "?S = ?E") + shows "sets (\\<^isub>M i\I. sigma (E i)) = sets (\\<^isub>M i\I. E i)" + (is "sets ?S = sets ?E") proof cases - assume "I = {}" then show ?thesis by (simp add: product_algebra_def) + assume "I = {}" then show ?thesis + by (simp add: product_algebra_def product_algebra_generator_def) next assume "I \ {}" interpret E: sigma_algebra "sigma (E i)" for i using E by (rule sigma_algebra_sigma) - have into_space[intro]: "\i x A. A \ sets (E i) \ x i \ A \ x i \ space (E i)" using E by auto - interpret G: sigma_algebra ?E - unfolding product_algebra_def using E - by (intro sigma_algebra_sigma) (auto dest: Pi_mem) - + unfolding product_algebra_def product_algebra_generator_def using E + by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) { fix A i assume "i \ I" and A: "A \ sets (E i)" then have "(\x. x i) -` A \ space ?E = (\\<^isub>E j\I. if j = i then A else \n. S j n) \ space ?E" - using isotone unfolding isoton_def product_algebra_def by (auto dest: Pi_mem) + using isotone unfolding isoton_def space_product_algebra + by (auto dest: Pi_mem) also have "\ = (\n. (\\<^isub>E j\I. if j = i then A else S j n))" - unfolding product_algebra_def + unfolding space_product_algebra apply simp apply (subst Pi_UN[OF `finite I`]) using isotone[THEN isoton_mono_le] apply simp apply (simp add: PiE_Int) apply (intro PiE_cong) using A sets_into by (auto intro!: into_space) - also have "\ \ sets ?E" unfolding product_algebra_def sets_sigma + also have "\ \ sets ?E" using sets_into `A \ sets (E i)` + unfolding sets_product_algebra sets_sigma by (intro sigma_sets.Union) (auto simp: image_subset_iff intro!: sigma_sets.Basic) finally have "(\x. x i) -` A \ space ?E \ sets ?E" . } then have proj: "\i. i\I \ (\x. x i) \ measurable ?E (sigma (E i))" using E by (subst G.measurable_iff_sigma) - (auto simp: product_algebra_def sets_sigma) - + (auto simp: sets_product_algebra sets_sigma) { fix A assume A: "\i. i \ I \ A i \ sets (sigma (E i))" with proj have basic: "\i. i \ I \ (\x. x i) -` (A i) \ space ?E \ sets ?E" unfolding measurable_def by simp @@ -1289,218 +1271,117 @@ using A E.sets_into_space `I \ {}` unfolding product_algebra_def by auto blast then have "Pi\<^isub>E I A \ sets ?E" using G.finite_INT[OF `finite I` `I \ {}` basic, of "\i. i"] by simp } - then have "sigma_sets (space ?E) (sets (product_algebra (\i. sigma (E i)) I)) \ sets ?E" - by (intro G.sigma_sets_subset) (auto simp add: sets_sigma product_algebra_def) + then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\i. sigma (E i)))) \ sets ?E" + by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def) then have subset: "sets ?S \ sets ?E" - by (simp add: sets_sigma product_algebra_def) - - have "sets ?S = sets ?E" + by (simp add: sets_sigma sets_product_algebra) + show "sets ?S = sets ?E" proof (intro set_eqI iffI) fix A assume "A \ sets ?E" then show "A \ sets ?S" - unfolding sets_sigma + unfolding sets_sigma sets_product_algebra proof induct case (Basic A) then show ?case - by (auto simp: sets_sigma product_algebra_def intro: sigma_sets.Basic) - qed (auto intro: sigma_sets.intros simp: product_algebra_def) + by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic) + qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def) next fix A assume "A \ sets ?S" then show "A \ sets ?E" using subset by auto qed - then show ?thesis - by (simp add: product_algebra_def sigma_def) +qed + +lemma product_algebraI[intro]: + "E \ (\ i\I. sets (M i)) \ Pi\<^isub>E I E \ sets (Pi\<^isub>M I M)" + using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI) + +lemma (in product_sigma_algebra) measurable_component_update: + assumes "x \ space (Pi\<^isub>M I M)" and "i \ I" + shows "(\v. x(i := v)) \ measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \ _") + unfolding product_algebra_def apply simp +proof (intro measurable_sigma) + let ?G = "product_algebra_generator (insert i I) M" + show "sets ?G \ Pow (space ?G)" using product_algebra_generator_into_space . + show "?f \ space (M i) \ space ?G" + using M.sets_into_space assms by auto + fix A assume "A \ sets ?G" + from product_algebraE[OF this] guess E . note E = this + then have "?f -` A \ space (M i) = E i \ ?f -` A \ space (M i) = {}" + using M.sets_into_space assms by auto + then show "?f -` A \ space (M i) \ sets (M i)" + using E by (auto intro!: product_algebraI) qed -lemma (in product_sigma_algebra) sigma_pair_algebra_sigma_eq: - "sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))) = - sigma (pair_algebra (product_algebra M I) (product_algebra M J))" - using M.sets_into_space - by (intro pair_sigma_algebra_sigma[of "\_. \\<^isub>E i\I. space (M i)", of _ "\_. \\<^isub>E i\J. space (M i)"]) - (auto simp: isoton_const product_algebra_def, blast+) - -lemma (in product_sigma_algebra) sigma_pair_algebra_product_singleton: - "sigma (pair_algebra (sigma (product_algebra M I)) (M i)) = - sigma (pair_algebra (product_algebra M I) (M i))" - using M.sets_into_space apply (subst M.sigma_eq[symmetric]) - by (intro pair_sigma_algebra_sigma[of "\_. \\<^isub>E i\I. space (M i)" _ "\_. space (M i)"]) - (auto simp: isoton_const product_algebra_def, blast+) - -lemma (in product_sigma_algebra) measurable_restrict: - assumes [simp]: "I \ J = {}" - shows "(\x. (restrict x I, restrict x J)) \ measurable - (sigma (product_algebra M (I \ J))) - (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))" - unfolding sigma_pair_algebra_sigma_eq using M.sets_into_space - by (intro measurable_sigma_sigma measurable_restrict_on_generating - pair_algebra_sets_into_space product_algebra_sets_into_space) - auto +lemma (in product_sigma_algebra) measurable_add_dim: + assumes "i \ I" + shows "(\(f, y). f(i := y)) \ measurable (Pi\<^isub>M I M \\<^isub>M M i) (Pi\<^isub>M (insert i I) M)" +proof - + let ?f = "(\(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M" + interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i" + unfolding pair_sigma_algebra_def + by (intro sigma_algebra_product_algebra sigma_algebras conjI) + have "?f \ measurable Ii.P (sigma ?G)" + proof (rule Ii.measurable_sigma) + show "sets ?G \ Pow (space ?G)" + using product_algebra_generator_into_space . + show "?f \ space (Pi\<^isub>M I M \\<^isub>M M i) \ space ?G" + by (auto simp: space_pair_measure) + next + fix A assume "A \ sets ?G" + then obtain F where "A = Pi\<^isub>E (insert i I) F" + and F: "\j. j \ I \ F j \ sets (M j)" "F i \ sets (M i)" + by (auto elim!: product_algebraE) + then have "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) = Pi\<^isub>E I F \ (F i)" + using sets_into_space `i \ I` + by (auto simp add: space_pair_measure) blast+ + then show "?f -` A \ space (Pi\<^isub>M I M \\<^isub>M M i) \ sets (Pi\<^isub>M I M \\<^isub>M M i)" + using F by (auto intro!: pair_measureI) + qed + then show ?thesis + by (simp add: product_algebra_def) +qed lemma (in product_sigma_algebra) measurable_merge: assumes [simp]: "I \ J = {}" - shows "(\(x, y). merge I x J y) \ measurable - (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J)))) - (sigma (product_algebra M (I \ J)))" - unfolding sigma_pair_algebra_sigma_eq using M.sets_into_space - by (intro measurable_sigma_sigma measurable_merge_on_generating - pair_algebra_sets_into_space product_algebra_sets_into_space) - auto - -lemma (in product_sigma_algebra) pair_product_product_vimage_algebra: - assumes [simp]: "I \ J = {}" - shows "sigma_algebra.vimage_algebra (sigma (product_algebra M (I \ J))) - (space (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))) (\(x,y). merge I x J y) = - (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))" - unfolding sigma_pair_algebra_sigma_eq using sets_into_space - by (intro vimage_algebra_sigma[OF bij_inv_restrict_merge[symmetric]] - pair_algebra_sets_into_space product_algebra_sets_into_space - measurable_merge_on_generating measurable_restrict_on_generating) - auto - -lemma (in product_sigma_algebra) measurable_restrict_iff: - assumes IJ[simp]: "I \ J = {}" - shows "f \ measurable (sigma (pair_algebra - (sigma (product_algebra M I)) (sigma (product_algebra M J)))) M' \ - (\x. f (restrict x I, restrict x J)) \ measurable (sigma (product_algebra M (I \ J))) M'" - using M.sets_into_space - apply (subst pair_product_product_vimage_algebra[OF IJ, symmetric]) - apply (subst sigma_pair_algebra_sigma_eq) - apply (subst sigma_algebra.measurable_vimage_iff_inv[OF _ - bij_inv_restrict_merge[symmetric]]) - apply (intro sigma_algebra_sigma product_algebra_sets_into_space) - by auto - -lemma (in product_sigma_algebra) measurable_merge_iff: - assumes IJ: "I \ J = {}" - shows "f \ measurable (sigma (product_algebra M (I \ J))) M' \ - (\(x, y). f (merge I x J y)) \ - measurable (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J)))) M'" - unfolding measurable_restrict_iff[OF IJ] - by (rule measurable_cong) (auto intro!: arg_cong[where f=f] simp: extensional_restrict) - -lemma (in product_sigma_algebra) measurable_component: - assumes "i \ I" and f: "f \ measurable (M i) M'" - shows "(\x. f (x i)) \ measurable (sigma (product_algebra M I)) M'" - (is "?f \ measurable ?P M'") + shows "(\(x, y). merge I x J y) \ measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" proof - - have "f \ (\x. x i) \ measurable ?P M'" - apply (rule measurable_comp[OF _ f]) - using measurable_up_sigma[of "product_algebra M I" "M i"] - using measurable_component_on_generator[OF `i \ I`] - by auto - then show "?f \ measurable ?P M'" by (simp add: comp_def) + let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M" + interpret P: sigma_algebra "?I \\<^isub>M ?J" + by (intro sigma_algebra_pair_measure product_algebra_into_space) + let ?f = "\(x, y). merge I x J y" + let ?G = "product_algebra_generator (I \ J) M" + have "?f \ measurable (?I \\<^isub>M ?J) (sigma ?G)" + proof (rule P.measurable_sigma) + fix A assume "A \ sets ?G" + from product_algebraE[OF this] + obtain E where E: "A = Pi\<^isub>E (I \ J) E" "E \ (\ i\I \ J. sets (M i))" . + then have *: "?f -` A \ space (?I \\<^isub>M ?J) = Pi\<^isub>E I E \ Pi\<^isub>E J E" + using sets_into_space `I \ J = {}` + by (auto simp add: space_pair_measure) blast+ + then show "?f -` A \ space (?I \\<^isub>M ?J) \ sets (?I \\<^isub>M ?J)" + using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI + simp: product_algebra_def) + qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure) + then show "?f \ measurable (?I \\<^isub>M ?J) (Pi\<^isub>M (I \ J) M)" + unfolding product_algebra_def[of "I \ J"] by simp qed -lemma (in product_sigma_algebra) singleton_vimage_algebra: - "sigma_algebra.vimage_algebra (sigma (product_algebra M {i})) (space (M i)) (\x. \j\{i}. x) = M i" - using sets_into_space - by (intro vimage_algebra_sigma[of "M i", unfolded M.sigma_eq, OF bij_inv_singleton[symmetric]] - product_algebra_sets_into_space measurable_singleton_on_generator measurable_component_on_generator) - auto - lemma (in product_sigma_algebra) measurable_component_singleton: - "(\x. f (x i)) \ measurable (sigma (product_algebra M {i})) M' \ - f \ measurable (M i) M'" - using sets_into_space - apply (subst singleton_vimage_algebra[symmetric]) - apply (subst sigma_algebra.measurable_vimage_iff_inv[OF _ bij_inv_singleton[symmetric]]) - by (auto intro!: sigma_algebra_sigma product_algebra_sets_into_space) - -lemma (in product_sigma_algebra) measurable_component_iff: - assumes "i \ I" and not_empty: "\i\I. space (M i) \ {}" - shows "(\x. f (x i)) \ measurable (sigma (product_algebra M I)) M' \ - f \ measurable (M i) M'" - (is "?f \ measurable ?P M' \ _") -proof - assume "f \ measurable (M i) M'" then show "?f \ measurable ?P M'" - by (rule measurable_component[OF `i \ I`]) -next - assume f: "?f \ measurable ?P M'" - def t \ "\i. SOME x. x \ space (M i)" - have t: "\i. i\I \ t i \ space (M i)" - unfolding t_def using not_empty by (rule_tac someI_ex) auto - have "?f \ (\x. (\j\I. if j = i then x else t j)) \ measurable (M i) M'" - (is "?f \ ?t \ measurable _ _") - proof (rule measurable_comp[OF _ f]) - have "?t \ measurable (M i) (product_algebra M I)" - proof (unfold measurable_def, intro CollectI conjI ballI) - from t show "?t \ space (M i) \ (space (product_algebra M I))" by auto - next - fix A assume A: "A \ sets (product_algebra M I)" - { fix E assume "E \ (\ i\I. sets (M i))" - then have "?t -` Pi\<^isub>E I E \ space (M i) = (if (\j\I-{i}. t j \ E j) then E i else {})" - using `i \ I` sets_into_space by (auto dest: Pi_mem[where B=E]) } - note * = this - with A `i \ I` show "?t -` A \ space (M i) \ sets (M i)" - by (auto elim!: product_algebraE simp del: vimage_Int) - qed - also have "\ \ measurable (M i) (sigma (product_algebra M I))" - using M.sets_into_space by (intro measurable_subset) (auto simp: product_algebra_def, blast) - finally show "?t \ measurable (M i) (sigma (product_algebra M I))" . - qed - then show "f \ measurable (M i) M'" unfolding comp_def using `i \ I` by simp -qed - -lemma (in product_sigma_algebra) measurable_singleton: - shows "f \ measurable (sigma (product_algebra M {i})) M' \ - (\x. f (\j\{i}. x)) \ measurable (M i) M'" - using sets_into_space unfolding measurable_component_singleton[symmetric] - by (auto intro!: measurable_cong arg_cong[where f=f] simp: fun_eq_iff extensional_def) - - -lemma (in pair_sigma_algebra) measurable_pair_split: - assumes "sigma_algebra M1'" "sigma_algebra M2'" - assumes f: "f \ measurable M1 M1'" and g: "g \ measurable M2 M2'" - shows "(\(x, y). (f x, g y)) \ measurable P (sigma (pair_algebra M1' M2'))" -proof (rule measurable_sigma) - interpret M1': sigma_algebra M1' by fact - interpret M2': sigma_algebra M2' by fact - interpret Q: pair_sigma_algebra M1' M2' by default - show "sets Q.E \ Pow (space Q.E)" - using M1'.sets_into_space M2'.sets_into_space by (auto simp: pair_algebra_def) - show "(\(x, y). (f x, g y)) \ space P \ space Q.E" - using f g unfolding measurable_def pair_algebra_def by auto - fix A assume "A \ sets Q.E" - then obtain X Y where A: "A = X \ Y" "X \ sets M1'" "Y \ sets M2'" - unfolding pair_algebra_def by auto - then have *: "(\(x, y). (f x, g y)) -` A \ space P = - (f -` X \ space M1) \ (g -` Y \ space M2)" - by (auto simp: pair_algebra_def) - then show "(\(x, y). (f x, g y)) -` A \ space P \ sets P" - using f g A unfolding measurable_def * - by (auto intro!: pair_algebraI in_sigma) -qed - -lemma (in product_sigma_algebra) measurable_add_dim: - assumes "i \ I" "finite I" - shows "(\(f, y). f(i := y)) \ measurable (sigma (pair_algebra (sigma (product_algebra M I)) (M i))) - (sigma (product_algebra M (insert i I)))" -proof (subst measurable_cong) - interpret I: finite_product_sigma_algebra M I by default fact - interpret i: finite_product_sigma_algebra M "{i}" by default auto - interpret Ii: pair_sigma_algebra I.P "M i" by default - interpret Ii': pair_sigma_algebra I.P i.P by default - have disj: "I \ {i} = {}" using `i \ I` by auto - have "(\(x, y). (id x, \_\{i}. y)) \ measurable Ii.P Ii'.P" - proof (intro Ii.measurable_pair_split I.measurable_ident) - show "(\y. \_\{i}. y) \ measurable (M i) i.P" - apply (rule measurable_singleton[THEN iffD1]) - using i.measurable_ident unfolding id_def . - qed default - from measurable_comp[OF this measurable_merge[OF disj]] - show "(\(x, y). merge I x {i} y) \ (\(x, y). (id x, \_\{i}. y)) - \ measurable (sigma (pair_algebra I.P (M i))) (sigma (product_algebra M (insert i I)))" - (is "?f \ _") by simp - fix x assume "x \ space Ii.P" - with assms show "(\(f, y). f(i := y)) x = ?f x" - by (cases x) (simp add: merge_def fun_eq_iff pair_algebra_def extensional_def) -qed + assumes "i \ I" shows "(\x. x i) \ measurable (Pi\<^isub>M I M) (M i)" +proof (unfold measurable_def, intro CollectI conjI ballI) + fix A assume "A \ sets (M i)" + then have "(\x. x i) -` A \ space (Pi\<^isub>M I M) = (\\<^isub>E j\I. if i = j then A else space (M j))" + using M.sets_into_space `i \ I` by (fastsimp dest: Pi_mem split: split_if_asm) + then show "(\x. x i) -` A \ space (Pi\<^isub>M I M) \ sets (Pi\<^isub>M I M)" + using `A \ sets (M i)` by (auto intro!: product_algebraI) +qed (insert `i \ I`, auto) locale product_sigma_finite = - fixes M :: "'i \ 'a algebra" and \ :: "'i \ 'a set \ pextreal" - assumes sigma_finite_measures: "\i. sigma_finite_measure (M i) (\ i)" + fixes M :: "'i \ ('a,'b) measure_space_scheme" + assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" -locale finite_product_sigma_finite = product_sigma_finite M \ for M :: "'i \ 'a algebra" and \ + - fixes I :: "'i set" assumes finite_index': "finite I" +locale finite_product_sigma_finite = product_sigma_finite M + for M :: "'i \ ('a,'b) measure_space_scheme" + + fixes I :: "'i set" assumes finite_index'[intro]: "finite I" -sublocale product_sigma_finite \ M: sigma_finite_measure "M i" "\ i" for i +sublocale product_sigma_finite \ M: sigma_finite_measure "M i" for i by (rule sigma_finite_measures) sublocale product_sigma_finite \ product_sigma_algebra @@ -1509,6 +1390,30 @@ sublocale finite_product_sigma_finite \ finite_product_sigma_algebra by default (fact finite_index') +lemma (in finite_product_sigma_finite) product_algebra_generator_measure: + assumes "Pi\<^isub>E I F \ sets G" + shows "measure G (Pi\<^isub>E I F) = (\i\I. M.\ i (F i))" +proof cases + assume ne: "\i\I. F i \ {}" + have "\i\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i" + by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) + (insert ne, auto simp: Pi_eq_iff) + then show ?thesis + unfolding product_algebra_generator_def by simp +next + assume empty: "\ (\j\I. F j \ {})" + then have "(\j\I. M.\ j (F j)) = 0" + by (auto simp: setprod_pextreal_0 intro!: bexI) + moreover + have "\j\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}" + by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) + (insert empty, auto simp: Pi_eq_empty_iff[symmetric]) + then have "(\j\I. M.\ j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" + by (auto simp: setprod_pextreal_0 intro!: bexI) + ultimately show ?thesis + unfolding product_algebra_generator_def by simp +qed + lemma (in finite_product_sigma_finite) sigma_finite_pairs: "\F::'i \ nat \ 'a set. (\i\I. range (F i) \ sets (M i)) \ @@ -1544,46 +1449,60 @@ lemma (in product_sigma_finite) product_measure_exists: assumes "finite I" - shows "\\. (\A\(\ i\I. sets (M i)). \ (Pi\<^isub>E I A) = (\i\I. \ i (A i))) \ - sigma_finite_measure (sigma (product_algebra M I)) \" + shows "\\. sigma_finite_measure (sigma (product_algebra_generator I M) \ measure := \ \) \ + (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i)))" using `finite I` proof induct - case empty then show ?case unfolding product_algebra_def - by (auto intro!: exI[of _ "\A. if A = {} then 0 else 1"] sigma_algebra_sigma - sigma_algebra.finite_additivity_sufficient - simp add: positive_def additive_def sets_sigma sigma_finite_measure_def - sigma_finite_measure_axioms_def image_constant) + case empty + interpret finite_product_sigma_finite M "{}" by default simp + let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ pextreal" + show ?case + proof (intro exI conjI ballI) + have "sigma_algebra (sigma G \measure := ?\\)" + by (rule sigma_algebra_cong) (simp_all add: product_algebra_def) + then have "measure_space (sigma G\measure := ?\\)" + by (rule finite_additivity_sufficient) + (simp_all add: positive_def additive_def sets_sigma + product_algebra_generator_def image_constant) + then show "sigma_finite_measure (sigma G\measure := ?\\)" + by (auto intro!: exI[of _ "\i. {\_. undefined}"] + simp: sigma_finite_measure_def sigma_finite_measure_axioms_def + product_algebra_generator_def) + qed auto next case (insert i I) - interpret finite_product_sigma_finite M \ I by default fact + interpret finite_product_sigma_finite M I by default fact have "finite (insert i I)" using `finite I` by auto - interpret I': finite_product_sigma_finite M \ "insert i I" by default fact + interpret I': finite_product_sigma_finite M "insert i I" by default fact from insert obtain \ where - prod: "\A. A \ (\ i\I. sets (M i)) \ \ (Pi\<^isub>E I A) = (\i\I. \ i (A i))" and - "sigma_finite_measure P \" by auto - interpret I: sigma_finite_measure P \ by fact - interpret P: pair_sigma_finite P \ "M i" "\ i" .. - + prod: "\A. A \ (\ i\I. sets (M i)) \ \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" and + "sigma_finite_measure (sigma G\ measure := \ \)" by auto + then interpret I: sigma_finite_measure "P\ measure := \\" unfolding product_algebra_def by simp + interpret P: pair_sigma_finite "P\ measure := \\" "M i" .. let ?h = "(\(f, y). f(i := y))" - let ?\ = "\A. P.pair_measure (?h -` A \ space P.P)" - have I': "sigma_algebra I'.P" by default - interpret I': measure_space "sigma (product_algebra M (insert i I))" ?\ - apply (rule P.measure_space_vimage[OF I']) - apply (rule measurable_add_dim[OF `i \ I` `finite I`]) - by simp + let ?\ = "\A. P.\ (?h -` A \ space P.P)" + have I': "sigma_algebra (I'.P\ measure := ?\ \)" + by (rule I'.sigma_algebra_cong) simp_all + interpret I'': measure_space "I'.P\ measure := ?\ \" + using measurable_add_dim[OF `i \ I`] + by (intro P.measure_space_vimage[OF I']) + (simp_all add: measurable_def pair_measure_def pair_measure_generator_def sets_sigma) show ?case proof (intro exI[of _ ?\] conjI ballI) + let "?m A" = "measure (Pi\<^isub>M I M\measure := \\ \\<^isub>M M i) (?h -` A \ space P.P)" { fix A assume A: "A \ (\ i\insert i I. sets (M i))" then have *: "?h -` Pi\<^isub>E (insert i I) A \ space P.P = Pi\<^isub>E I A \ A i" - using `i \ I` M.sets_into_space by (auto simp: pair_algebra_def) blast - show "?\ (Pi\<^isub>E (insert i I) A) = (\i\insert i I. \ i (A i))" + using `i \ I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast + show "?m (Pi\<^isub>E (insert i I) A) = (\i\insert i I. M.\ i (A i))" unfolding * using A apply (subst P.pair_measure_times) using A apply fastsimp using A apply fastsimp using `i \ I` `finite I` prod[of A] A by (auto simp: ac_simps) } note product = this - show "sigma_finite_measure I'.P ?\" - proof + have *: "sigma I'.G\ measure := ?\ \ = I'.P\ measure := ?\ \" + by (simp add: product_algebra_def) + show "sigma_finite_measure (sigma I'.G\ measure := ?\ \)" + proof (unfold *, default, simp) from I'.sigma_finite_pairs guess F :: "'i \ nat \ 'a set" .. then have F: "\j. j \ insert i I \ range (F j) \ sets (M j)" "(\k. \\<^isub>E j \ insert i I. F j k) \ space I'.G" @@ -1591,12 +1510,12 @@ by blast+ let "?F k" = "\\<^isub>E j \ insert i I. F j k" show "\F::nat \ ('i \ 'a) set. range F \ sets I'.P \ - (\i. F i) = space I'.P \ (\i. ?\ (F i) \ \)" + (\i. F i) = (\\<^isub>E i\insert i I. space (M i)) \ (\i. ?m (F i) \ \)" proof (intro exI[of _ ?F] conjI allI) show "range ?F \ sets I'.P" using F(1) by auto next from F(2)[THEN isotonD(2)] - show "(\i. ?F i) = space I'.P" by simp + show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp next fix j show "?\ (?F j) \ \" @@ -1607,62 +1526,43 @@ qed qed -definition (in finite_product_sigma_finite) - measure :: "('i \ 'a) set \ pextreal" where - "measure = (SOME \. - (\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. \ i (A i))) \ - sigma_finite_measure P \)" - -sublocale finite_product_sigma_finite \ sigma_finite_measure P measure -proof - - show "sigma_finite_measure P measure" - unfolding measure_def - by (rule someI2_ex[OF product_measure_exists[OF finite_index]]) auto -qed +sublocale finite_product_sigma_finite \ sigma_finite_measure P + unfolding product_algebra_def + using product_measure_exists[OF finite_index] + by (rule someI2_ex) auto lemma (in finite_product_sigma_finite) measure_times: assumes "\i. i \ I \ A i \ sets (M i)" - shows "measure (Pi\<^isub>E I A) = (\i\I. \ i (A i))" -proof - - note ex = product_measure_exists[OF finite_index] - show ?thesis - unfolding measure_def - proof (rule someI2_ex[OF ex], elim conjE) - fix \ assume *: "\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. \ i (A i))" + shows "\ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" + unfolding product_algebra_def + using product_measure_exists[OF finite_index] + proof (rule someI2_ex, elim conjE) + fix \ assume *: "\A\\ i\I. sets (M i). \ (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" have "Pi\<^isub>E I A = Pi\<^isub>E I (\i\I. A i)" by (auto dest: Pi_mem) then have "\ (Pi\<^isub>E I A) = \ (Pi\<^isub>E I (\i\I. A i))" by simp - also have "\ = (\i\I. \ i ((\i\I. A i) i))" using assms * by auto - finally show "\ (Pi\<^isub>E I A) = (\i\I. \ i (A i))" by simp + also have "\ = (\i\I. M.\ i ((\i\I. A i) i))" using assms * by auto + finally show "measure (sigma G\measure := \\) (Pi\<^isub>E I A) = (\i\I. M.\ i (A i))" + by (simp add: sigma_def) qed -qed - -abbreviation (in product_sigma_finite) - "product_measure I \ finite_product_sigma_finite.measure M \ I" - -abbreviation (in product_sigma_finite) - "product_positive_integral I \ - measure_space.positive_integral (sigma (product_algebra M I)) (product_measure I)" - -abbreviation (in product_sigma_finite) - "product_integral I \ - measure_space.integral (sigma (product_algebra M I)) (product_measure I)" - -abbreviation (in product_sigma_finite) - "product_integrable I \ - measure_space.integrable (sigma (product_algebra M I)) (product_measure I)" lemma (in product_sigma_finite) product_measure_empty[simp]: - "product_measure {} {\x. undefined} = 1" + "measure (Pi\<^isub>M {} M) {\x. undefined} = 1" proof - - interpret finite_product_sigma_finite M \ "{}" by default auto + interpret finite_product_sigma_finite M "{}" by default auto from measure_times[of "\x. {}"] show ?thesis by simp qed +lemma (in finite_product_sigma_algebra) P_empty: + assumes "I = {}" + shows "space P = {\k. undefined}" "sets P = { {}, {\k. undefined} }" + unfolding product_algebra_def product_algebra_generator_def `I = {}` + by (simp_all add: sigma_def image_constant) + lemma (in product_sigma_finite) positive_integral_empty: - "product_positive_integral {} f = f (\k. undefined)" + "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\k. undefined)" proof - - interpret finite_product_sigma_finite M \ "{}" by default (fact finite.emptyI) - have "\A. measure (Pi\<^isub>E {} A) = 1" + interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) + have "\A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1" using assms by (subst measure_times) auto then show ?thesis unfolding positive_integral_def simple_function_def simple_integral_def_raw @@ -1676,18 +1576,15 @@ lemma (in product_sigma_finite) measure_fold: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - assumes A: "A \ sets (sigma (product_algebra M (I \ J)))" - shows "pair_sigma_finite.pair_measure - (sigma (product_algebra M I)) (product_measure I) - (sigma (product_algebra M J)) (product_measure J) - ((\(x,y). merge I x J y) -` A \ space (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))) = - product_measure (I \ J) A" + assumes A: "A \ sets (Pi\<^isub>M (I \ J) M)" + shows "measure (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) ((\(x,y). merge I x J y) -` A \ space (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)) = + measure (Pi\<^isub>M (I \ J) M) A" proof - - interpret I: finite_product_sigma_finite M \ I by default fact - interpret J: finite_product_sigma_finite M \ J by default fact + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M \ "I \ J" by default fact - interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite I.P J.P by default let ?g = "\(x,y). merge I x J y" let "?X S" = "?g -` S \ space P.P" from IJ.sigma_finite_pairs obtain F where @@ -1696,111 +1593,125 @@ "\k. \i\I\J. \ i (F i k) \ \" by auto let ?F = "\k. \\<^isub>E i\I \ J. F i k" - show "P.pair_measure (?X A) = IJ.measure A" - proof (rule measure_unique_Int_stable[OF _ _ _ _ _ _ _ _ A]) - show "Int_stable IJ.G" by (simp add: PiE_Int Int_stable_def product_algebra_def) auto - show "range ?F \ sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def) + show "P.\ (?X A) = IJ.\ A" + proof (rule measure_unique_Int_stable[where X=A]) + show "A \ sets (sigma IJ.G)" + using A unfolding product_algebra_def by auto + show "Int_stable IJ.G" + by (simp add: PiE_Int Int_stable_def product_algebra_def + product_algebra_generator_def) + auto + show "range ?F \ sets IJ.G" using F + by (simp add: image_subset_iff product_algebra_def + product_algebra_generator_def) show "?F \ space IJ.G " using F(2) by simp - have "sigma_algebra IJ.P" by default - then show "measure_space IJ.P (\A. P.pair_measure (?X A))" - apply (rule P.measure_space_vimage) - apply (rule measurable_merge[OF `I \ J = {}`]) - by simp - show "measure_space IJ.P IJ.measure" by fact + have "measure_space IJ.P" by fact + also have "IJ.P = \ space = space IJ.G, sets = sets (sigma IJ.G), measure = IJ.\ \" + by (simp add: product_algebra_def) + finally show "measure_space \" . + let ?P = "\ space = space IJ.G, sets = sets (sigma IJ.G), + measure = \A. P.\ (?X A)\" + have *: "?P = (sigma IJ.G \ measure := \A. P.\ (?X A) \)" + by auto + have "sigma_algebra (sigma IJ.G \ measure := \A. P.\ (?X A) \)" + by (rule IJ.sigma_algebra_cong) (auto simp: product_algebra_def) + then show "measure_space ?P" unfolding * + using measurable_merge[OF `I \ J = {}`] + by (auto intro!: P.measure_space_vimage simp add: product_algebra_def) next fix A assume "A \ sets IJ.G" then obtain F where A[simp]: "A = Pi\<^isub>E (I \ J) F" and F: "\i. i \ I \ J \ F i \ sets (M i)" - by (auto simp: product_algebra_def) + by (auto simp: product_algebra_generator_def) then have "?X A = (Pi\<^isub>E I F \ Pi\<^isub>E J F)" - using sets_into_space by (auto simp: space_pair_algebra) blast+ - then have "P.pair_measure (?X A) = (\i\I. \ i (F i)) * (\i\J. \ i (F i))" + using sets_into_space by (auto simp: space_pair_measure) blast+ + then have "P.\ (?X A) = (\i\I. \ i (F i)) * (\i\J. \ i (F i))" using `finite J` `finite I` F by (simp add: P.pair_measure_times I.measure_times J.measure_times) also have "\ = (\i\I \ J. \ i (F i))" using `finite J` `finite I` `I \ J = {}` by (simp add: setprod_Un_one) - also have "\ = IJ.measure A" + also have "\ = IJ.\ A" using `finite J` `finite I` F unfolding A by (intro IJ.measure_times[symmetric]) auto - finally show "P.pair_measure (?X A) = IJ.measure A" . + finally show "P.\ (?X A) = IJ.\ A" . next fix k have k: "\i. i \ I \ J \ F i k \ sets (M i)" using F by auto then have "?X (?F k) = (\\<^isub>E i\I. F i k) \ (\\<^isub>E i\J. F i k)" - using sets_into_space by (auto simp: space_pair_algebra) blast+ - with k have "P.pair_measure (?X (?F k)) = (\i\I. \ i (F i k)) * (\i\J. \ i (F i k))" + using sets_into_space by (auto simp: space_pair_measure) blast+ + with k have "P.\ (?X (?F k)) = (\i\I. \ i (F i k)) * (\i\J. \ i (F i k))" by (simp add: P.pair_measure_times I.measure_times J.measure_times) - then show "P.pair_measure (?X (?F k)) \ \" + then show "P.\ (?X (?F k)) \ \" using `finite I` F by (simp add: setprod_\) - qed simp + qed qed lemma (in product_sigma_finite) product_positive_integral_fold: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - and f: "f \ borel_measurable (sigma (product_algebra M (I \ J)))" - shows "product_positive_integral (I \ J) f = - product_positive_integral I (\x. product_positive_integral J (\y. f (merge I x J y)))" + and f: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" + shows "integral\<^isup>P (Pi\<^isub>M (I \ J) M) f = + (\\<^isup>+ x. (\\<^isup>+ y. f (merge I x J y) \(Pi\<^isub>M J M)) \(Pi\<^isub>M I M))" proof - - interpret I: finite_product_sigma_finite M \ I by default fact - interpret J: finite_product_sigma_finite M \ J by default fact + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M \ "I \ J" by default fact - interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite I.P J.P by default have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable P.P" - unfolding case_prod_distrib measurable_merge_iff[OF IJ, symmetric] using f . + using measurable_comp[OF measurable_merge[OF IJ] f] by (simp add: comp_def) show ?thesis unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] - apply (subst IJ.positive_integral_cong_measure[symmetric]) - apply (rule measure_fold[OF IJ fin]) - apply assumption proof (rule P.positive_integral_vimage) show "sigma_algebra IJ.P" by default show "(\(x, y). merge I x J y) \ measurable P.P IJ.P" by (rule measurable_merge[OF IJ]) - show "f \ borel_measurable IJ.P" using f . + show "f \ borel_measurable IJ.P" using f by simp + next + fix A assume "A \ sets (Pi\<^isub>M (I \ J) M)" + then show "IJ.\ A = P.\ ((\(x,y). merge I x J y) -` A \ space P.P)" + using measure_fold[OF IJ fin] by simp qed qed lemma (in product_sigma_finite) product_positive_integral_singleton: assumes f: "f \ borel_measurable (M i)" - shows "product_positive_integral {i} (\x. f (x i)) = M.positive_integral i f" + shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\x. f (x i)) = integral\<^isup>P (M i) f" proof - - interpret I: finite_product_sigma_finite M \ "{i}" by default simp - have T: "(\x. x i) \ measurable (sigma (product_algebra M {i})) (M i)" - using measurable_component_singleton[of "\x. x" i] - measurable_ident[unfolded id_def] by auto - show "I.positive_integral (\x. f (x i)) = M.positive_integral i f" - unfolding I.positive_integral_vimage[OF sigma_algebras T f, symmetric] - proof (rule positive_integral_cong_measure) - fix A let ?P = "(\x. x i) -` A \ space (sigma (product_algebra M {i}))" + interpret I: finite_product_sigma_finite M "{i}" by default simp + show ?thesis + proof (rule I.positive_integral_vimage[symmetric]) + show "sigma_algebra (M i)" by (rule sigma_algebras) + show "(\x. x i) \ measurable (Pi\<^isub>M {i} M) (M i)" by (rule measurable_component_singleton) auto + show "f \ borel_measurable (M i)" by fact + next + fix A let ?P = "(\x. x i) -` A \ space I.P" assume A: "A \ sets (M i)" then have *: "?P = {i} \\<^isub>E A" using sets_into_space by auto - show "product_measure {i} ?P = \ i A" unfolding * + show "M.\ i A = I.\ ?P" unfolding * using A I.measure_times[of "\_. A"] by auto qed qed lemma (in product_sigma_finite) product_positive_integral_insert: assumes [simp]: "finite I" "i \ I" - and f: "f \ borel_measurable (sigma (product_algebra M (insert i I)))" - shows "product_positive_integral (insert i I) f - = product_positive_integral I (\x. M.positive_integral i (\y. f (x(i:=y))))" + and f: "f \ borel_measurable (Pi\<^isub>M (insert i I) M)" + shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\\<^isup>+ x. (\\<^isup>+ y. f (x(i := y)) \(M i)) \(Pi\<^isub>M I M))" proof - - interpret I: finite_product_sigma_finite M \ I by default auto - interpret i: finite_product_sigma_finite M \ "{i}" by default auto + interpret I: finite_product_sigma_finite M I by default auto + interpret i: finite_product_sigma_finite M "{i}" by default auto interpret P: pair_sigma_algebra I.P i.P .. - have IJ: "I \ {i} = {}" by auto + have IJ: "I \ {i} = {}" and insert: "I \ {i} = insert i I" + using f by auto show ?thesis - unfolding product_positive_integral_fold[OF IJ, simplified, OF f] + unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f] proof (rule I.positive_integral_cong, subst product_positive_integral_singleton) fix x assume x: "x \ space I.P" let "?f y" = "f (restrict (x(i := y)) (insert i I))" have f'_eq: "\y. ?f y = f (x(i := y))" using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) - note fP = f[unfolded measurable_merge_iff[OF IJ, simplified]] - show "?f \ borel_measurable (M i)" - using P.measurable_pair_image_snd[OF fP x] - unfolding measurable_singleton f'_eq by (simp add: f'_eq) - show "M.positive_integral i ?f = M.positive_integral i (\y. f (x(i := y)))" + show "?f \ borel_measurable (M i)" unfolding f'_eq + using measurable_comp[OF measurable_component_update f] x `i \ I` + by (simp add: comp_def) + show "integral\<^isup>P (M i) ?f = \\<^isup>+ y. f (x(i:=y)) \M i" unfolding f'_eq by simp qed qed @@ -1808,24 +1719,21 @@ lemma (in product_sigma_finite) product_positive_integral_setprod: fixes f :: "'i \ 'a \ pextreal" assumes "finite I" and borel: "\i. i \ I \ f i \ borel_measurable (M i)" - shows "product_positive_integral I (\x. (\i\I. f i (x i))) = - (\i\I. M.positive_integral i (f i))" + shows "(\\<^isup>+ x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>P (M i) (f i))" using assms proof induct case empty - interpret finite_product_sigma_finite M \ "{}" by default auto + interpret finite_product_sigma_finite M "{}" by default auto then show ?case by simp next case (insert i I) note `finite I`[intro, simp] - interpret I: finite_product_sigma_finite M \ I by default auto + interpret I: finite_product_sigma_finite M I by default auto have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using insert by (auto intro!: setprod_cong) - have prod: "\J. J \ insert i I \ - (\x. (\i\J. f i (x i))) \ borel_measurable (sigma (product_algebra M J))" + have prod: "\J. J \ insert i I \ (\x. (\i\J. f i (x i))) \ borel_measurable (Pi\<^isub>M J M)" using sets_into_space insert - by (intro sigma_algebra.borel_measurable_pextreal_setprod - sigma_algebra_sigma product_algebra_sets_into_space - measurable_component) + by (intro sigma_algebra.borel_measurable_pextreal_setprod sigma_algebra_product_algebra + measurable_comp[OF measurable_component_singleton, unfolded comp_def]) auto show ?case by (simp add: product_positive_integral_insert[OF insert(1,2) prod]) @@ -1834,67 +1742,51 @@ lemma (in product_sigma_finite) product_integral_singleton: assumes f: "f \ borel_measurable (M i)" - shows "product_integral {i} (\x. f (x i)) = M.integral i f" + shows "(\x. f (x i) \Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f" proof - - interpret I: finite_product_sigma_finite M \ "{i}" by default simp + interpret I: finite_product_sigma_finite M "{i}" by default simp have *: "(\x. Real (f x)) \ borel_measurable (M i)" "(\x. Real (- f x)) \ borel_measurable (M i)" using assms by auto show ?thesis - unfolding I.integral_def integral_def - unfolding *[THEN product_positive_integral_singleton] .. + unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. qed lemma (in product_sigma_finite) product_integral_fold: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - and f: "measure_space.integrable (sigma (product_algebra M (I \ J))) (product_measure (I \ J)) f" - shows "product_integral (I \ J) f = - product_integral I (\x. product_integral J (\y. f (merge I x J y)))" + and f: "integrable (Pi\<^isub>M (I \ J) M) f" + shows "integral\<^isup>L (Pi\<^isub>M (I \ J) M) f = (\x. (\y. f (merge I x J y) \Pi\<^isub>M J M) \Pi\<^isub>M I M)" proof - - interpret I: finite_product_sigma_finite M \ I by default fact - interpret J: finite_product_sigma_finite M \ J by default fact + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M \ "I \ J" by default fact - interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default - let ?f = "\(x,y). f (merge I x J y)" - have f_borel: "f \ borel_measurable IJ.P" - "(\x. Real (f x)) \ borel_measurable IJ.P" - "(\x. Real (- f x)) \ borel_measurable IJ.P" - using f unfolding integrable_def by auto - have f_restrict: "(\x. f (restrict x (I \ J))) \ borel_measurable IJ.P" - by (rule measurable_cong[THEN iffD2, OF _ f_borel(1)]) - (auto intro!: arg_cong[where f=f] simp: extensional_restrict) - then have f'_borel: - "(\x. Real (?f x)) \ borel_measurable P.P" - "(\x. Real (- ?f x)) \ borel_measurable P.P" - unfolding measurable_restrict_iff[OF IJ] - by simp_all - have PI: - "P.positive_integral (\x. Real (?f x)) = IJ.positive_integral (\x. Real (f x))" - "P.positive_integral (\x. Real (- ?f x)) = IJ.positive_integral (\x. Real (- f x))" - using f'_borel[THEN P.positive_integral_fst_measurable(2)] - using f_borel(2,3)[THEN product_positive_integral_fold[OF assms(1-3)]] - by simp_all - have "P.integrable ?f" using `IJ.integrable f` - unfolding P.integrable_def IJ.integrable_def - unfolding measurable_restrict_iff[OF IJ] - using f_restrict PI by simp_all + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite I.P J.P by default + let ?M = "\(x, y). merge I x J y" + let ?f = "\x. f (?M x)" show ?thesis - unfolding P.integrable_fst_measurable(2)[OF `P.integrable ?f`, simplified] - unfolding IJ.integral_def P.integral_def - unfolding PI by simp + proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) + have 1: "sigma_algebra IJ.P" by default + have 2: "?M \ measurable P.P IJ.P" using measurable_merge[OF IJ] . + have 3: "\A. A \ sets IJ.P \ IJ.\ A = P.\ (?M -` A \ space P.P)" + by (rule measure_fold[OF IJ fin, symmetric]) + have 4: "integrable (Pi\<^isub>M (I \ J) M) f" by fact + show "integrable P.P ?f" + by (rule P.integral_vimage[where f=f, OF 1 2 3 4]) + show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" + by (rule P.integral_vimage[where f=f, OF 1 2 3 4]) + qed qed lemma (in product_sigma_finite) product_integral_insert: assumes [simp]: "finite I" "i \ I" - and f: "measure_space.integrable (sigma (product_algebra M (insert i I))) (product_measure (insert i I)) f" - shows "product_integral (insert i I) f - = product_integral I (\x. M.integral i (\y. f (x(i:=y))))" + and f: "integrable (Pi\<^isub>M (insert i I) M) f" + shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^isub>M I M)" proof - - interpret I: finite_product_sigma_finite M \ I by default auto - interpret I': finite_product_sigma_finite M \ "insert i I" by default auto - interpret i: finite_product_sigma_finite M \ "{i}" by default auto - interpret P: pair_sigma_algebra I.P i.P .. + interpret I: finite_product_sigma_finite M I by default auto + interpret I': finite_product_sigma_finite M "insert i I" by default auto + interpret i: finite_product_sigma_finite M "{i}" by default auto + interpret P: pair_sigma_finite I.P i.P .. have IJ: "I \ {i} = {}" by auto show ?thesis unfolding product_integral_fold[OF IJ, simplified, OF f] @@ -1903,39 +1795,40 @@ let "?f y" = "f (restrict (x(i := y)) (insert i I))" have f'_eq: "\y. ?f y = f (x(i := y))" using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) - have "f \ borel_measurable I'.P" using f unfolding I'.integrable_def by auto - note fP = this[unfolded measurable_merge_iff[OF IJ, simplified]] + have f: "f \ borel_measurable I'.P" using f unfolding integrable_def by auto show "?f \ borel_measurable (M i)" - using P.measurable_pair_image_snd[OF fP x] - unfolding measurable_singleton f'_eq by (simp add: f'_eq) - show "M.integral i ?f = M.integral i (\y. f (x(i := y)))" + unfolding measurable_cong[OF f'_eq] + using measurable_comp[OF measurable_component_update f] x `i \ I` + by (simp add: comp_def) + show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\y. f (x(i := y)))" unfolding f'_eq by simp qed qed lemma (in product_sigma_finite) product_integrable_setprod: fixes f :: "'i \ 'a \ real" - assumes [simp]: "finite I" and integrable: "\i. i \ I \ M.integrable i (f i)" - shows "product_integrable I (\x. (\i\I. f i (x i)))" (is "product_integrable I ?f") + assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "integrable (Pi\<^isub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") proof - - interpret finite_product_sigma_finite M \ I by default fact + interpret finite_product_sigma_finite M I by default fact have f: "\i. i \ I \ f i \ borel_measurable (M i)" - using integrable unfolding M.integrable_def by auto + using integrable unfolding integrable_def by auto then have borel: "?f \ borel_measurable P" - by (intro borel_measurable_setprod measurable_component) auto - moreover have "integrable (\x. \\i\I. f i (x i)\)" + using measurable_comp[OF measurable_component_singleton f] + by (auto intro!: borel_measurable_setprod simp: comp_def) + moreover have "integrable (Pi\<^isub>M I M) (\x. \\i\I. f i (x i)\)" proof (unfold integrable_def, intro conjI) show "(\x. abs (?f x)) \ borel_measurable P" using borel by auto - have "positive_integral (\x. Real (abs (?f x))) = positive_integral (\x. \i\I. Real (abs (f i (x i))))" + have "(\\<^isup>+x. Real (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. Real (abs (f i (x i)))) \P)" by (simp add: Real_setprod abs_setprod) - also have "\ = (\i\I. M.positive_integral i (\x. Real (abs (f i x))))" + also have "\ = (\i\I. (\\<^isup>+x. Real (abs (f i x)) \M i))" using f by (subst product_positive_integral_setprod) auto also have "\ < \" using integrable[THEN M.integrable_abs] - unfolding pextreal_less_\ setprod_\ M.integrable_def by simp - finally show "positive_integral (\x. Real (abs (?f x))) \ \" by auto - show "positive_integral (\x. Real (- abs (?f x))) \ \" by simp + unfolding pextreal_less_\ setprod_\ integrable_def by simp + finally show "(\\<^isup>+x. Real (abs (?f x)) \P) \ \" by auto + show "(\\<^isup>+x. Real (- abs (?f x)) \P) \ \" by simp qed ultimately show ?thesis by (rule integrable_abs_iff[THEN iffD1]) @@ -1943,8 +1836,8 @@ lemma (in product_sigma_finite) product_integral_setprod: fixes f :: "'i \ 'a \ real" - assumes "finite I" "I \ {}" and integrable: "\i. i \ I \ M.integrable i (f i)" - shows "product_integral I (\x. (\i\I. f i (x i))) = (\i\I. M.integral i (f i))" + assumes "finite I" "I \ {}" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "(\x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>L (M i) (f i))" using assms proof (induct rule: finite_ne_induct) case (singleton i) then show ?case by (simp add: product_integral_singleton integrable_def) @@ -1952,9 +1845,9 @@ case (insert i I) then have iI: "finite (insert i I)" by auto then have prod: "\J. J \ insert i I \ - product_integrable J (\x. (\i\J. f i (x i)))" + integrable (Pi\<^isub>M J M) (\x. (\i\J. f i (x i)))" by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) - interpret I: finite_product_sigma_finite M \ I by default fact + interpret I: finite_product_sigma_finite M I by default fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using `i \ I` by (auto intro!: setprod_cong) show ?case @@ -1964,9 +1857,9 @@ section "Products on finite spaces" -lemma sigma_sets_pair_algebra_finite: +lemma sigma_sets_pair_measure_generator_finite: assumes "finite A" and "finite B" - shows "sigma_sets (A \ B) ((\(x,y). x \ y) ` (Pow A \ Pow B)) = Pow (A \ B)" + shows "sigma_sets (A \ B) { a \ b | a b. a \ Pow A \ b \ Pow B} = Pow (A \ B)" (is "sigma_sets ?prod ?sets = _") proof safe have fin: "finite (A \ B)" using assms by (rule finite_cartesian_product) @@ -1978,7 +1871,7 @@ next case (insert a x) hence "{a} \ sigma_sets ?prod ?sets" - by (auto simp: pair_algebra_def intro!: sigma_sets.Basic) + by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic) moreover have "x \ sigma_sets ?prod ?sets" using insert by auto ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) qed @@ -1989,26 +1882,28 @@ show "a \ A" and "b \ B" by auto qed -locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 M2 +locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 + for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" sublocale pair_finite_sigma_algebra \ pair_sigma_algebra by default -lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra[simp]: - shows "P = (| space = space M1 \ space M2, sets = Pow (space M1 \ space M2) |)" +lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra: + shows "P = \ space = space M1 \ space M2, sets = Pow (space M1 \ space M2), \ = algebra.more P \" proof - - show ?thesis using M1.finite_space M2.finite_space - by (simp add: sigma_def space_pair_algebra sets_pair_algebra - sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow) + show ?thesis + using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space] + by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def) qed sublocale pair_finite_sigma_algebra \ finite_sigma_algebra P -proof - show "finite (space P)" "sets P = Pow (space P)" - using M1.finite_space M2.finite_space by auto -qed + apply default + using M1.finite_space M2.finite_space + apply (subst finite_pair_sigma_algebra) apply simp + apply (subst (1 2) finite_pair_sigma_algebra) apply simp + done -locale pair_finite_space = M1: finite_measure_space M1 \1 + M2: finite_measure_space M2 \2 - for M1 \1 M2 \2 +locale pair_finite_space = M1: finite_measure_space M1 + M2: finite_measure_space M2 + for M1 M2 sublocale pair_finite_space \ pair_finite_sigma_algebra by default @@ -2016,19 +1911,11 @@ sublocale pair_finite_space \ pair_sigma_finite by default -lemma (in pair_finite_space) finite_pair_sigma_algebra[simp]: - shows "P = (| space = space M1 \ space M2, sets = Pow (space M1 \ space M2) |)" -proof - - show ?thesis using M1.finite_space M2.finite_space - by (simp add: sigma_def space_pair_algebra sets_pair_algebra - sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow) -qed - lemma (in pair_finite_space) pair_measure_Pair[simp]: assumes "a \ space M1" "b \ space M2" - shows "pair_measure {(a, b)} = \1 {a} * \2 {b}" + shows "\ {(a, b)} = M1.\ {a} * M2.\ {b}" proof - - have "pair_measure ({a}\{b}) = \1 {a} * \2 {b}" + have "\ ({a}\{b}) = M1.\ {a} * M2.\ {b}" using M1.sets_eq_Pow M2.sets_eq_Pow assms by (subst pair_measure_times) auto then show ?thesis by simp @@ -2036,15 +1923,10 @@ lemma (in pair_finite_space) pair_measure_singleton[simp]: assumes "x \ space M1 \ space M2" - shows "pair_measure {x} = \1 {fst x} * \2 {snd x}" + shows "\ {x} = M1.\ {fst x} * M2.\ {snd x}" using pair_measure_Pair assms by (cases x) auto -sublocale pair_finite_space \ finite_measure_space P pair_measure - by default auto - -lemma (in pair_finite_space) finite_measure_space_finite_prod_measure_alterantive: - "finite_measure_space \space = space M1 \ space M2, sets = Pow (space M1 \ space M2)\ pair_measure" - unfolding finite_pair_sigma_algebra[symmetric] - by default +sublocale pair_finite_space \ finite_measure_space P + by default (auto simp: space_pair_measure) end \ No newline at end of file diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Radon_Nikodym.thy Wed Feb 02 12:34:45 2011 +0100 @@ -11,7 +11,7 @@ qed auto lemma (in sigma_finite_measure) Ex_finite_integrable_function: - shows "\h\borel_measurable M. positive_integral h \ \ \ (\x\space M. 0 < h x \ h x < \)" + shows "\h\borel_measurable M. integral\<^isup>P M h \ \ \ (\x\space M. 0 < h x \ h x < \)" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and @@ -42,7 +42,7 @@ proof (safe intro!: bexI[of _ ?h] del: notI) have "\i. A i \ sets M" using range by fastsimp+ - then have "positive_integral ?h = (\\<^isub>\ i. n i * \ (A i))" + then have "integral\<^isup>P M ?h = (\\<^isub>\ i. n i * \ (A i))" by (simp add: positive_integral_psuminf positive_integral_cmult_indicator) also have "\ \ (\\<^isub>\ i. Real ((1 / 2)^Suc i))" proof (rule psuminf_le) @@ -56,7 +56,7 @@ qed also have "\ = Real 1" by (rule suminf_imp_psuminf, rule power_half_series, auto) - finally show "positive_integral ?h \ \" by auto + finally show "integral\<^isup>P M ?h \ \" by auto next fix x assume "x \ space M" then obtain i where "x \ A i" using space[symmetric] by auto @@ -75,46 +75,47 @@ "absolutely_continuous \ = (\N\null_sets. \ N = (0 :: pextreal))" lemma (in sigma_finite_measure) absolutely_continuous_AE: - assumes "measure_space M \" "absolutely_continuous \" "AE x. P x" - shows "measure_space.almost_everywhere M \ P" + assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M" + and "absolutely_continuous (measure M')" "AE x. P x" + shows "measure_space.almost_everywhere M' P" proof - - interpret \: measure_space M \ by fact + interpret \: measure_space M' by fact from `AE x. P x` obtain N where N: "N \ null_sets" and "{x\space M. \ P x} \ N" unfolding almost_everywhere_def by auto show "\.almost_everywhere P" proof (rule \.AE_I') - show "{x\space M. \ P x} \ N" by fact - from `absolutely_continuous \` show "N \ \.null_sets" + show "{x\space M'. \ P x} \ N" by simp fact + from `absolutely_continuous (measure M')` show "N \ \.null_sets" using N unfolding absolutely_continuous_def by auto qed qed lemma (in finite_measure_space) absolutely_continuousI: - assumes "finite_measure_space M \" + assumes "finite_measure_space (M\ measure := \\)" (is "finite_measure_space ?\") assumes v: "\x. \ x \ space M ; \ {x} = 0 \ \ \ {x} = 0" shows "absolutely_continuous \" proof (unfold absolutely_continuous_def sets_eq_Pow, safe) fix N assume "\ N = 0" "N \ space M" - interpret v: finite_measure_space M \ by fact - have "\ N = \ (\x\N. {x})" by simp - also have "\ = (\x\N. \ {x})" + interpret v: finite_measure_space ?\ by fact + have "\ N = measure ?\ (\x\N. {x})" by simp + also have "\ = (\x\N. measure ?\ {x})" proof (rule v.measure_finitely_additive''[symmetric]) show "finite N" using `N \ space M` finite_space by (auto intro: finite_subset) show "disjoint_family_on (\i. {i}) N" unfolding disjoint_family_on_def by auto - fix x assume "x \ N" thus "{x} \ sets M" using `N \ space M` sets_eq_Pow by auto + fix x assume "x \ N" thus "{x} \ sets ?\" using `N \ space M` sets_eq_Pow by auto qed also have "\ = 0" proof (safe intro!: setsum_0') fix x assume "x \ N" hence "\ {x} \ \ N" using sets_eq_Pow `N \ space M` by (auto intro!: measure_mono) hence "\ {x} = 0" using `\ N = 0` by simp - thus "\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto + thus "measure ?\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto qed - finally show "\ N = 0" . + finally show "\ N = 0" by simp qed lemma (in measure_space) density_is_absolutely_continuous: - assumes "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x)" + assumes "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x \M)" shows "absolutely_continuous \" using assms unfolding absolutely_continuous_def by (simp add: positive_integral_null_set) @@ -123,13 +124,13 @@ lemma (in finite_measure) Radon_Nikodym_aux_epsilon: fixes e :: real assumes "0 < e" - assumes "finite_measure M s" - shows "\A\sets M. real (\ (space M)) - real (s (space M)) \ - real (\ A) - real (s A) \ - (\B\sets M. B \ A \ - e < real (\ B) - real (s B))" + assumes "finite_measure (M\measure := \\)" + shows "\A\sets M. real (\ (space M)) - real (\ (space M)) \ + real (\ A) - real (\ A) \ + (\B\sets M. B \ A \ - e < real (\ B) - real (\ B))" proof - - let "?d A" = "real (\ A) - real (s A)" - interpret M': finite_measure M s by fact + let "?d A" = "real (\ A) - real (\ A)" + interpret M': finite_measure "M\measure := \\" by fact let "?A A" = "if (\B\sets M. B \ space M - A \ -e < ?d B) then {} else (SOME B. B \ sets M \ B \ space M - A \ ?d B \ -e)" @@ -216,21 +217,24 @@ qed lemma (in finite_measure) Radon_Nikodym_aux: - assumes "finite_measure M s" - shows "\A\sets M. real (\ (space M)) - real (s (space M)) \ - real (\ A) - real (s A) \ - (\B\sets M. B \ A \ 0 \ real (\ B) - real (s B))" + assumes "finite_measure (M\measure := \\)" (is "finite_measure ?M'") + shows "\A\sets M. real (\ (space M)) - real (\ (space M)) \ + real (\ A) - real (\ A) \ + (\B\sets M. B \ A \ 0 \ real (\ B) - real (\ B))" proof - - let "?d A" = "real (\ A) - real (s A)" + let "?d A" = "real (\ A) - real (\ A)" let "?P A B n" = "A \ sets M \ A \ B \ ?d B \ ?d A \ (\C\sets M. C \ A \ - 1 / real (Suc n) < ?d C)" - interpret M': finite_measure M s by fact + interpret M': finite_measure ?M' where + "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \" by fact auto let "?r S" = "restricted_space S" { fix S n assume S: "S \ sets M" hence **: "\X. X \ op \ S ` sets M \ X \ sets M \ X \ S" by auto - from M'.restricted_finite_measure[of S] restricted_finite_measure[of S] S - have "finite_measure (?r S) \" "0 < 1 / real (Suc n)" - "finite_measure (?r S) s" by auto + have [simp]: "(restricted_space S\measure := \\) = M'.restricted_space S" + by (cases M) simp + from M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S + have "finite_measure (?r S)" "0 < 1 / real (Suc n)" + "finite_measure (?r S\measure := \\)" by auto from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. hence "?P X S n" proof (simp add: **, safe) @@ -287,12 +291,14 @@ qed lemma (in finite_measure) Radon_Nikodym_finite_measure: - assumes "finite_measure M \" + assumes "finite_measure (M\ measure := \\)" (is "finite_measure ?M'") assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x)" + shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M)" proof - - interpret M': finite_measure M \ using assms(1) . - def G \ "{g \ borel_measurable M. \A\sets M. (\\<^isup>+x. g x * indicator A x) \ \ A}" + interpret M': finite_measure ?M' + where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \" + using assms(1) by auto + def G \ "{g \ borel_measurable M. \A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ \ A}" have "(\x. 0) \ G" unfolding G_def by auto hence "G \ {}" by auto { fix f g assume f: "f \ G" and g: "g \ G" @@ -308,16 +314,16 @@ have "\x. x \ space M \ max (g x) (f x) * indicator A x = g x * indicator (?A \ A) x + f x * indicator ((space M - ?A) \ A) x" by (auto simp: indicator_def max_def) - hence "(\\<^isup>+x. max (g x) (f x) * indicator A x) = - (\\<^isup>+x. g x * indicator (?A \ A) x) + - (\\<^isup>+x. f x * indicator ((space M - ?A) \ A) x)" + hence "(\\<^isup>+x. max (g x) (f x) * indicator A x \M) = + (\\<^isup>+x. g x * indicator (?A \ A) x \M) + + (\\<^isup>+x. f x * indicator ((space M - ?A) \ A) x \M)" using f g sets unfolding G_def by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator) also have "\ \ \ (?A \ A) + \ ((space M - ?A) \ A)" using f g sets unfolding G_def by (auto intro!: add_mono) also have "\ = \ A" using M'.measure_additive[OF sets] union by auto - finally show "(\\<^isup>+x. max (g x) (f x) * indicator A x) \ \ A" . + finally show "(\\<^isup>+x. max (g x) (f x) * indicator A x \M) \ \ A" . qed } note max_in_G = this { fix f g assume "f \ g" and f: "\i. f i \ G" @@ -331,30 +337,30 @@ hence "\i. (\x. f i x * indicator A x) \ borel_measurable M" using f_borel by (auto intro!: borel_measurable_indicator) from positive_integral_isoton[OF isoton_indicator[OF `f \ g`] this] - have SUP: "(\\<^isup>+x. g x * indicator A x) = - (SUP i. (\\<^isup>+x. f i x * indicator A x))" + have SUP: "(\\<^isup>+x. g x * indicator A x \M) = + (SUP i. (\\<^isup>+x. f i x * indicator A x \M))" unfolding isoton_def by simp - show "(\\<^isup>+x. g x * indicator A x) \ \ A" unfolding SUP + show "(\\<^isup>+x. g x * indicator A x \M) \ \ A" unfolding SUP using f `A \ sets M` unfolding G_def by (auto intro!: SUP_leI) qed } note SUP_in_G = this - let ?y = "SUP g : G. positive_integral g" + let ?y = "SUP g : G. integral\<^isup>P M g" have "?y \ \ (space M)" unfolding G_def proof (safe intro!: SUP_leI) - fix g assume "\A\sets M. (\\<^isup>+x. g x * indicator A x) \ \ A" - from this[THEN bspec, OF top] show "positive_integral g \ \ (space M)" + fix g assume "\A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ \ A" + from this[THEN bspec, OF top] show "integral\<^isup>P M g \ \ (space M)" by (simp cong: positive_integral_cong) qed hence "?y \ \" using M'.finite_measure_of_space by auto from SUPR_countable_SUPR[OF this `G \ {}`] guess ys .. note ys = this - hence "\n. \g. g\G \ positive_integral g = ys n" + hence "\n. \g. g\G \ integral\<^isup>P M g = ys n" proof safe - fix n assume "range ys \ positive_integral ` G" - hence "ys n \ positive_integral ` G" by auto - thus "\g. g\G \ positive_integral g = ys n" by auto + fix n assume "range ys \ integral\<^isup>P M ` G" + hence "ys n \ integral\<^isup>P M ` G" by auto + thus "\g. g\G \ integral\<^isup>P M g = ys n" by auto qed - from choice[OF this] obtain gs where "\i. gs i \ G" "\n. positive_integral (gs n) = ys n" by auto - hence y_eq: "?y = (SUP i. positive_integral (gs i))" using ys by auto + from choice[OF this] obtain gs where "\i. gs i \ G" "\n. integral\<^isup>P M (gs n) = ys n" by auto + hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto let "?g i x" = "Max ((\n. gs n x) ` {..i})" def f \ "SUP i. ?g i" have gs_not_empty: "\i. (\n. gs n x) ` {..i} \ {}" by auto @@ -372,53 +378,53 @@ hence isoton_g: "?g \ f" by (simp add: isoton_def le_fun_def f_def) from SUP_in_G[OF this g_in_G] have "f \ G" . hence [simp, intro]: "f \ borel_measurable M" unfolding G_def by auto - have "(\i. positive_integral (?g i)) \ positive_integral f" + have "(\i. integral\<^isup>P M (?g i)) \ integral\<^isup>P M f" using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def) - hence "positive_integral f = (SUP i. positive_integral (?g i))" + hence "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding isoton_def by simp also have "\ = ?y" proof (rule antisym) - show "(SUP i. positive_integral (?g i)) \ ?y" + show "(SUP i. integral\<^isup>P M (?g i)) \ ?y" using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def) - show "?y \ (SUP i. positive_integral (?g i))" unfolding y_eq + show "?y \ (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq by (auto intro!: SUP_mono positive_integral_mono Max_ge) qed - finally have int_f_eq_y: "positive_integral f = ?y" . - let "?t A" = "\ A - (\\<^isup>+x. f x * indicator A x)" - have "finite_measure M ?t" - proof - show "?t {} = 0" by simp - show "?t (space M) \ \" using M'.finite_measure by simp - show "countably_additive M ?t" unfolding countably_additive_def - proof safe - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" - have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x)) - = (\\<^isup>+x. (\\<^isub>\n. f x * indicator (A n) x))" - using `range A \ sets M` - by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator) - also have "\ = (\\<^isup>+x. f x * indicator (\n. A n) x)" - apply (rule positive_integral_cong) - apply (subst psuminf_cmult_right) - unfolding psuminf_indicator[OF `disjoint_family A`] .. - finally have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x)) - = (\\<^isup>+x. f x * indicator (\n. A n) x)" . - moreover have "(\\<^isub>\n. \ (A n)) = \ (\n. A n)" - using M'.measure_countably_additive A by (simp add: comp_def) - moreover have "\i. (\\<^isup>+x. f x * indicator (A i) x) \ \ (A i)" - using A `f \ G` unfolding G_def by auto - moreover have v_fin: "\ (\i. A i) \ \" using M'.finite_measure A by (simp add: countable_UN) - moreover { - have "(\\<^isup>+x. f x * indicator (\i. A i) x) \ \ (\i. A i)" - using A `f \ G` unfolding G_def by (auto simp: countable_UN) - also have "\ (\i. A i) < \" using v_fin by (simp add: pextreal_less_\) - finally have "(\\<^isup>+x. f x * indicator (\i. A i) x) \ \" - by (simp add: pextreal_less_\) } - ultimately - show "(\\<^isub>\ n. ?t (A n)) = ?t (\i. A i)" - apply (subst psuminf_minus) by simp_all - qed + finally have int_f_eq_y: "integral\<^isup>P M f = ?y" . + let "?t A" = "\ A - (\\<^isup>+x. f x * indicator A x \M)" + let ?M = "M\ measure := ?t\" + interpret M: sigma_algebra ?M + by (intro sigma_algebra_cong) auto + have fmM: "finite_measure ?M" + proof (default, simp_all add: countably_additive_def, safe) + fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" + have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x \M)) + = (\\<^isup>+x. (\\<^isub>\n. f x * indicator (A n) x) \M)" + using `range A \ sets M` + by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator) + also have "\ = (\\<^isup>+x. f x * indicator (\n. A n) x \M)" + apply (rule positive_integral_cong) + apply (subst psuminf_cmult_right) + unfolding psuminf_indicator[OF `disjoint_family A`] .. + finally have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x \M)) + = (\\<^isup>+x. f x * indicator (\n. A n) x \M)" . + moreover have "(\\<^isub>\n. \ (A n)) = \ (\n. A n)" + using M'.measure_countably_additive A by (simp add: comp_def) + moreover have "\i. (\\<^isup>+x. f x * indicator (A i) x \M) \ \ (A i)" + using A `f \ G` unfolding G_def by auto + moreover have v_fin: "\ (\i. A i) \ \" using M'.finite_measure A by (simp add: countable_UN) + moreover { + have "(\\<^isup>+x. f x * indicator (\i. A i) x \M) \ \ (\i. A i)" + using A `f \ G` unfolding G_def by (auto simp: countable_UN) + also have "\ (\i. A i) < \" using v_fin by (simp add: pextreal_less_\) + finally have "(\\<^isup>+x. f x * indicator (\i. A i) x \M) \ \" + by (simp add: pextreal_less_\) } + ultimately + show "(\\<^isub>\ n. ?t (A n)) = ?t (\i. A i)" + apply (subst psuminf_minus) by simp_all qed - then interpret M: finite_measure M ?t . + then interpret M: finite_measure ?M + where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t" + by (simp_all add: fmM) have ac: "absolutely_continuous ?t" using `absolutely_continuous \` unfolding absolutely_continuous_def by auto have upper_bound: "\A\sets M. ?t A \ 0" proof (rule ccontr) @@ -433,23 +439,21 @@ hence pos_M: "0 < \ (space M)" using ac top unfolding absolutely_continuous_def by auto moreover - have "(\\<^isup>+x. f x * indicator (space M) x) \ \ (space M)" + have "(\\<^isup>+x. f x * indicator (space M) x \M) \ \ (space M)" using `f \ G` unfolding G_def by auto - hence "(\\<^isup>+x. f x * indicator (space M) x) \ \" + hence "(\\<^isup>+x. f x * indicator (space M) x \M) \ \" using M'.finite_measure_of_space by auto moreover def b \ "?t (space M) / \ (space M) / 2" ultimately have b: "b \ 0" "b \ \" using M'.finite_measure_of_space by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space) - have "finite_measure M (\A. b * \ A)" (is "finite_measure M ?b") - proof - show "?b {} = 0" by simp - show "?b (space M) \ \" using finite_measure_of_space b by auto - show "countably_additive M ?b" - unfolding countably_additive_def psuminf_cmult_right - using measure_countably_additive by auto - qed + let ?Mb = "?M\measure := \A. b * \ A\" + interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto + have "finite_measure ?Mb" + by default + (insert finite_measure_of_space b measure_countably_additive, + auto simp: psuminf_cmult_right countably_additive_def) from M.Radon_Nikodym_aux[OF this] obtain A0 where "A0 \ sets M" and space_less_A0: "real (?t (space M)) - real (b * \ (space M)) \ real (?t A0) - real (b * \ A0)" and @@ -462,30 +466,30 @@ let "?f0 x" = "f x + b * indicator A0 x" { fix A assume A: "A \ sets M" hence "A \ A0 \ sets M" using `A0 \ sets M` by auto - have "(\\<^isup>+x. ?f0 x * indicator A x) = - (\\<^isup>+x. f x * indicator A x + b * indicator (A \ A0) x)" + have "(\\<^isup>+x. ?f0 x * indicator A x \M) = + (\\<^isup>+x. f x * indicator A x + b * indicator (A \ A0) x \M)" by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith) - hence "(\\<^isup>+x. ?f0 x * indicator A x) = - (\\<^isup>+x. f x * indicator A x) + b * \ (A \ A0)" + hence "(\\<^isup>+x. ?f0 x * indicator A x \M) = + (\\<^isup>+x. f x * indicator A x \M) + b * \ (A \ A0)" using `A0 \ sets M` `A \ A0 \ sets M` A by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) } note f0_eq = this { fix A assume A: "A \ sets M" hence "A \ A0 \ sets M" using `A0 \ sets M` by auto - have f_le_v: "(\\<^isup>+x. f x * indicator A x) \ \ A" + have f_le_v: "(\\<^isup>+x. f x * indicator A x \M) \ \ A" using `f \ G` A unfolding G_def by auto note f0_eq[OF A] - also have "(\\<^isup>+x. f x * indicator A x) + b * \ (A \ A0) \ - (\\<^isup>+x. f x * indicator A x) + ?t (A \ A0)" + also have "(\\<^isup>+x. f x * indicator A x \M) + b * \ (A \ A0) \ + (\\<^isup>+x. f x * indicator A x \M) + ?t (A \ A0)" using bM_le_t[OF `A \ A0 \ sets M`] `A \ sets M` `A0 \ sets M` by (auto intro!: add_left_mono) - also have "\ \ (\\<^isup>+x. f x * indicator A x) + ?t A" + also have "\ \ (\\<^isup>+x. f x * indicator A x \M) + ?t A" using M.measure_mono[simplified, OF _ `A \ A0 \ sets M` `A \ sets M`] by (auto intro!: add_left_mono) also have "\ \ \ A" using f_le_v M'.finite_measure[simplified, OF `A \ sets M`] - by (cases "(\\<^isup>+x. f x * indicator A x)", cases "\ A", auto) - finally have "(\\<^isup>+x. ?f0 x * indicator A x) \ \ A" . } + by (cases "(\\<^isup>+x. f x * indicator A x \M)", cases "\ A", auto) + finally have "(\\<^isup>+x. ?f0 x * indicator A x \M) \ \ A" . } hence "?f0 \ G" using `A0 \ sets M` unfolding G_def by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times) have real: "?t (space M) \ \" "?t A0 \ \" @@ -494,7 +498,7 @@ finite_measure[of A0] M.finite_measure[of A0] finite_measure_of_space M.finite_measure_of_space by auto - have int_f_finite: "positive_integral f \ \" + have int_f_finite: "integral\<^isup>P M f \ \" using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff by (auto cong: positive_integral_cong) have "?t (space M) > b * \ (space M)" unfolding b_def @@ -514,22 +518,22 @@ using `A0 \ sets M` by auto hence "0 < b * \ A0" using b by auto from int_f_finite this - have "?y + 0 < positive_integral f + b * \ A0" unfolding int_f_eq_y + have "?y + 0 < integral\<^isup>P M f + b * \ A0" unfolding int_f_eq_y by (rule pextreal_less_add) - also have "\ = positive_integral ?f0" using f0_eq[OF top] `A0 \ sets M` sets_into_space + also have "\ = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \ sets M` sets_into_space by (simp cong: positive_integral_cong) - finally have "?y < positive_integral ?f0" by simp - moreover from `?f0 \ G` have "positive_integral ?f0 \ ?y" by (auto intro!: le_SUPI) + finally have "?y < integral\<^isup>P M ?f0" by simp + moreover from `?f0 \ G` have "integral\<^isup>P M ?f0 \ ?y" by (auto intro!: le_SUPI) ultimately show False by auto qed show ?thesis proof (safe intro!: bexI[of _ f]) fix A assume "A\sets M" - show "\ A = (\\<^isup>+x. f x * indicator A x)" + show "\ A = (\\<^isup>+x. f x * indicator A x \M)" proof (rule antisym) - show "(\\<^isup>+x. f x * indicator A x) \ \ A" + show "(\\<^isup>+x. f x * indicator A x \M) \ \ A" using `f \ G` `A \ sets M` unfolding G_def by auto - show "\ A \ (\\<^isup>+x. f x * indicator A x)" + show "\ A \ (\\<^isup>+x. f x * indicator A x \M)" using upper_bound[THEN bspec, OF `A \ sets M`] by (simp add: pextreal_zero_le_diff) qed @@ -537,13 +541,15 @@ qed lemma (in finite_measure) split_space_into_finite_sets_and_rest: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" (is "measure_space ?N") assumes ac: "absolutely_continuous \" shows "\\0\sets M. \\::nat\'a set. disjoint_family \ \ range \ \ sets M \ \0 = space M - (\i. \ i) \ (\A\sets M. A \ \0 \ (\ A = 0 \ \ A = 0) \ (\ A > 0 \ \ A = \)) \ (\i. \ (\ i) \ \)" proof - - interpret v: measure_space M \ by fact + interpret v: measure_space ?N + where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \" + by fact auto let ?Q = "{Q\sets M. \ Q \ \}" let ?a = "SUP Q:?Q. \ Q" have "{} \ ?Q" using v.empty_measure by auto @@ -667,11 +673,13 @@ qed lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" (is "measure_space ?N") assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x)" + shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M)" proof - - interpret v: measure_space M \ by fact + interpret v: measure_space ?N + where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \" + by fact auto from split_space_into_finite_sets_and_rest[OF assms] obtain Q0 and Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" @@ -680,39 +688,38 @@ and Q_fin: "\i. \ (Q i) \ \" by force from Q have Q_sets: "\i. Q i \ sets M" by auto have "\i. \f. f\borel_measurable M \ (\A\sets M. - \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x))" + \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" proof fix i have indicator_eq: "\f x A. (f x :: pextreal) * indicator (Q i \ A) x * indicator (Q i) x = (f x * indicator (Q i) x) * indicator A x" unfolding indicator_def by auto - have fm: "finite_measure (restricted_space (Q i)) \" - (is "finite_measure ?R \") by (rule restricted_finite_measure[OF Q_sets[of i]]) + have fm: "finite_measure (restricted_space (Q i))" + (is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]]) then interpret R: finite_measure ?R . - have fmv: "finite_measure ?R \" + have fmv: "finite_measure (restricted_space (Q i) \ measure := \\)" (is "finite_measure ?Q") unfolding finite_measure_def finite_measure_axioms_def proof - show "measure_space ?R \" + show "measure_space ?Q" using v.restricted_measure_space Q_sets[of i] by auto - show "\ (space ?R) \ \" - using Q_fin by simp + show "measure ?Q (space ?Q) \ \" using Q_fin by simp qed have "R.absolutely_continuous \" using `absolutely_continuous \` `Q i \ sets M` by (auto simp: R.absolutely_continuous_def absolutely_continuous_def) - from finite_measure.Radon_Nikodym_finite_measure[OF fm fmv this] + from R.Radon_Nikodym_finite_measure[OF fmv this] obtain f where f: "(\x. f x * indicator (Q i) x) \ borel_measurable M" - and f_int: "\A. A\sets M \ \ (Q i \ A) = (\\<^isup>+x. (f x * indicator (Q i) x) * indicator A x)" + and f_int: "\A. A\sets M \ \ (Q i \ A) = (\\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \M)" unfolding Bex_def borel_measurable_restricted[OF `Q i \ sets M`] positive_integral_restricted[OF `Q i \ sets M`] by (auto simp: indicator_eq) then show "\f. f\borel_measurable M \ (\A\sets M. - \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x))" + \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" by (fastsimp intro!: exI[of _ "\x. f x * indicator (Q i) x"] positive_integral_cong simp: indicator_def) qed from choice[OF this] obtain f where borel: "\i. f i \ borel_measurable M" and f: "\A i. A \ sets M \ - \ (Q i \ A) = (\\<^isup>+x. f i x * indicator (Q i \ A) x)" + \ (Q i \ A) = (\\<^isup>+x. f i x * indicator (Q i \ A) x \M)" by auto let "?f x" = "(\\<^isub>\ i. f i x * indicator (Q i) x) + \ * indicator Q0 x" @@ -728,7 +735,7 @@ f i x * indicator (Q i \ A) x" "\x i. (indicator A x * indicator Q0 x :: pextreal) = indicator (Q0 \ A) x" by (auto simp: indicator_def) - have "(\\<^isup>+x. ?f x * indicator A x) = + have "(\\<^isup>+x. ?f x * indicator A x \M) = (\\<^isub>\ i. \ (Q i \ A)) + \ * \ (Q0 \ A)" unfolding f[OF `A \ sets M`] apply (simp del: pextreal_times(2) add: field_simps *) @@ -755,27 +762,29 @@ using Q_sets `A \ sets M` Q0(1) by (auto intro!: countable_UN) moreover have "((\i. Q i) \ A) \ (Q0 \ A) = A" "((\i. Q i) \ A) \ (Q0 \ A) = {}" using `A \ sets M` sets_into_space Q0 by auto - ultimately show "\ A = (\\<^isup>+x. ?f x * indicator A x)" + ultimately show "\ A = (\\<^isup>+x. ?f x * indicator A x \M)" using v.measure_additive[simplified, of "(\i. Q i) \ A" "Q0 \ A"] by simp qed qed lemma (in sigma_finite_measure) Radon_Nikodym: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" (is "measure_space ?N") assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x)" + shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M)" proof - from Ex_finite_integrable_function - obtain h where finite: "positive_integral h \ \" and + obtain h where finite: "integral\<^isup>P M h \ \" and borel: "h \ borel_measurable M" and pos: "\x. x \ space M \ 0 < h x" and "\x. x \ space M \ h x < \" by auto - let "?T A" = "(\\<^isup>+x. h x * indicator A x)" + let "?T A" = "(\\<^isup>+x. h x * indicator A x \M)" + let ?MT = "M\ measure := ?T \" from measure_space_density[OF borel] finite - interpret T: finite_measure M ?T + interpret T: finite_measure ?MT + where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T" unfolding finite_measure_def finite_measure_axioms_def - by (simp cong: positive_integral_cong) + by (simp_all cong: positive_integral_cong) have "\N. N \ sets M \ {x \ space M. h x \ 0 \ indicator N x \ (0::pextreal)} = N" using sets_into_space pos by (force simp: indicator_def) then have "T.absolutely_continuous \" using assms(2) borel @@ -783,7 +792,8 @@ by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff) from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this] obtain f where f_borel: "f \ borel_measurable M" and - fT: "\A. A \ sets M \ \ A = T.positive_integral (\x. f x * indicator A x)" by auto + fT: "\A. A \ sets M \ \ A = (\\<^isup>+ x. f x * indicator A x \?MT)" + by (auto simp: measurable_def) show ?thesis proof (safe intro!: bexI[of _ "\x. h x * f x"]) show "(\x. h x * f x) \ borel_measurable M" @@ -792,7 +802,7 @@ then have "(\x. f x * indicator A x) \ borel_measurable M" using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator) from positive_integral_translated_density[OF borel this] - show "\ A = (\\<^isup>+x. h x * f x * indicator A x)" + show "\ A = (\\<^isup>+x. h x * f x * indicator A x \M)" unfolding fT[OF `A \ sets M`] by (simp add: field_simps) qed qed @@ -801,8 +811,8 @@ lemma (in measure_space) finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" - and fin: "positive_integral f < \" - shows "(\A\sets M. (\\<^isup>+x. f x * indicator A x) = (\\<^isup>+x. g x * indicator A x)) + and fin: "integral\<^isup>P M f < \" + shows "(\A\sets M. (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. g x * indicator A x \M)) \ (AE x. f x = g x)" (is "(\A\sets M. ?P f A = ?P g A) \ _") proof (intro iffI ballI) @@ -812,18 +822,18 @@ next assume eq: "\A\sets M. ?P f A = ?P g A" from this[THEN bspec, OF top] fin - have g_fin: "positive_integral g < \" by (simp cong: positive_integral_cong) + have g_fin: "integral\<^isup>P M g < \" by (simp cong: positive_integral_cong) { fix f g assume borel: "f \ borel_measurable M" "g \ borel_measurable M" - and g_fin: "positive_integral g < \" and eq: "\A\sets M. ?P f A = ?P g A" + and g_fin: "integral\<^isup>P M g < \" and eq: "\A\sets M. ?P f A = ?P g A" let ?N = "{x\space M. g x < f x}" have N: "?N \ sets M" using borel by simp - have "?P (\x. (f x - g x)) ?N = (\\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x)" + have "?P (\x. (f x - g x)) ?N = (\\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \M)" by (auto intro!: positive_integral_cong simp: indicator_def) also have "\ = ?P f ?N - ?P g ?N" proof (rule positive_integral_diff) show "(\x. f x * indicator ?N x) \ borel_measurable M" "(\x. g x * indicator ?N x) \ borel_measurable M" using borel N by auto - have "?P g ?N \ positive_integral g" + have "?P g ?N \ integral\<^isup>P M g" by (auto intro!: positive_integral_mono simp: indicator_def) then show "?P g ?N \ \" using g_fin by auto fix x assume "x \ space M" @@ -848,17 +858,17 @@ lemma (in finite_measure) density_unique_finite_measure: assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" - assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x) = (\\<^isup>+x. f' x * indicator A x)" + assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. f' x * indicator A x \M)" (is "\A. A \ sets M \ ?P f A = ?P f' A") shows "AE x. f x = f' x" proof - let "?\ A" = "?P f A" and "?\' A" = "?P f' A" let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x" - interpret M: measure_space M ?\ - using borel(1) by (rule measure_space_density) + interpret M: measure_space "M\ measure := ?\\" + using borel(1) by (rule measure_space_density) simp have ac: "absolutely_continuous ?\" using f by (rule density_is_absolutely_continuous) - from split_space_into_finite_sets_and_rest[OF `measure_space M ?\` ac] + from split_space_into_finite_sets_and_rest[OF `measure_space (M\ measure := ?\\)` ac] obtain Q0 and Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" and Q0: "Q0 \ sets M" "Q0 = space M - (\i. Q i)" @@ -876,13 +886,13 @@ have 2: "AE x. ?f Q0 x = ?f' Q0 x" proof (rule AE_I') { fix f :: "'a \ pextreal" assume borel: "f \ borel_measurable M" - and eq: "\A. A \ sets M \ ?\ A = (\\<^isup>+x. f x * indicator A x)" + and eq: "\A. A \ sets M \ ?\ A = (\\<^isup>+x. f x * indicator A x \M)" let "?A i" = "Q0 \ {x \ space M. f x < of_nat i}" have "(\i. ?A i) \ null_sets" proof (rule null_sets_UN) fix i have "?A i \ sets M" using borel Q0(1) by auto - have "?\ (?A i) \ (\\<^isup>+x. of_nat i * indicator (?A i) x)" + have "?\ (?A i) \ (\\<^isup>+x. of_nat i * indicator (?A i) x \M)" unfolding eq[OF `?A i \ sets M`] by (auto intro!: positive_integral_mono simp: indicator_def) also have "\ = of_nat i * \ (?A i)" @@ -912,63 +922,72 @@ lemma (in sigma_finite_measure) density_unique: assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" - assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x) = (\\<^isup>+x. f' x * indicator A x)" + assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. f' x * indicator A x \M)" (is "\A. A \ sets M \ ?P f A = ?P f' A") shows "AE x. f x = f' x" proof - obtain h where h_borel: "h \ borel_measurable M" - and fin: "positive_integral h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" + and fin: "integral\<^isup>P M h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" using Ex_finite_integrable_function by auto - interpret h: measure_space M "\A. (\\<^isup>+x. h x * indicator A x)" - using h_borel by (rule measure_space_density) - interpret h: finite_measure M "\A. (\\<^isup>+x. h x * indicator A x)" + interpret h: measure_space "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" + using h_borel by (rule measure_space_density) simp + interpret h: finite_measure "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" by default (simp cong: positive_integral_cong add: fin) - interpret f: measure_space M "\A. (\\<^isup>+x. f x * indicator A x)" - using borel(1) by (rule measure_space_density) - interpret f': measure_space M "\A. (\\<^isup>+x. f' x * indicator A x)" - using borel(2) by (rule measure_space_density) + let ?fM = "M\measure := \A. (\\<^isup>+x. f x * indicator A x \M)\" + interpret f: measure_space ?fM + using borel(1) by (rule measure_space_density) simp + let ?f'M = "M\measure := \A. (\\<^isup>+x. f' x * indicator A x \M)\" + interpret f': measure_space ?f'M + using borel(2) by (rule measure_space_density) simp { fix A assume "A \ sets M" then have " {x \ space M. h x \ 0 \ indicator A x \ (0::pextreal)} = A" using pos sets_into_space by (force simp: indicator_def) - then have "(\\<^isup>+x. h x * indicator A x) = 0 \ A \ null_sets" + then have "(\\<^isup>+x. h x * indicator A x \M) = 0 \ A \ null_sets" using h_borel `A \ sets M` by (simp add: positive_integral_0_iff) } note h_null_sets = this { fix A assume "A \ sets M" - have "(\\<^isup>+x. h x * (f x * indicator A x)) = - f.positive_integral (\x. h x * indicator A x)" + have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * indicator A x \?fM)" + using `A \ sets M` h_borel borel + by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) + also have "\ = (\\<^isup>+x. h x * indicator A x \?f'M)" + by (rule f'.positive_integral_cong_measure) (simp_all add: f) + also have "\ = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" using `A \ sets M` h_borel borel by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) - also have "\ = f'.positive_integral (\x. h x * indicator A x)" - by (rule f'.positive_integral_cong_measure) (rule f) - also have "\ = (\\<^isup>+x. h x * (f' x * indicator A x))" - using `A \ sets M` h_borel borel - by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) - finally have "(\\<^isup>+x. h x * (f x * indicator A x)) = (\\<^isup>+x. h x * (f' x * indicator A x))" . } + finally have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" . } then have "h.almost_everywhere (\x. f x = f' x)" using h_borel borel - by (intro h.density_unique_finite_measure[OF borel]) - (simp add: positive_integral_translated_density) + apply (intro h.density_unique_finite_measure) + apply (simp add: measurable_def) + apply (simp add: measurable_def) + by (simp add: positive_integral_translated_density) then show "AE x. f x = f' x" unfolding h.almost_everywhere_def almost_everywhere_def by (auto simp add: h_null_sets) qed lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: - assumes \: "measure_space M \" and f: "f \ borel_measurable M" - and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x)" - shows "sigma_finite_measure M \ \ (AE x. f x \ \)" + assumes \: "measure_space (M\measure := \\)" (is "measure_space ?N") + and f: "f \ borel_measurable M" + and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x \M)" + shows "sigma_finite_measure (M\measure := \\) \ (AE x. f x \ \)" proof - assume "sigma_finite_measure M \" - then interpret \: sigma_finite_measure M \ . + assume "sigma_finite_measure ?N" + then interpret \: sigma_finite_measure ?N + where "borel_measurable ?N = borel_measurable M" and "measure ?N = \" + and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def) from \.Ex_finite_integrable_function obtain h where - h: "h \ borel_measurable M" "\.positive_integral h \ \" + h: "h \ borel_measurable M" "integral\<^isup>P ?N h \ \" and fin: "\x. x \ space M \ 0 < h x \ h x < \" by auto have "AE x. f x * h x \ \" proof (rule AE_I') - have "\.positive_integral h = (\\<^isup>+x. f x * h x)" - by (simp add: \.positive_integral_cong_measure[symmetric, OF eq[symmetric]]) - (intro positive_integral_translated_density f h) - then have "(\\<^isup>+x. f x * h x) \ \" + have "integral\<^isup>P ?N h = (\\<^isup>+x. f x * h x \M)" + apply (subst \.positive_integral_cong_measure[symmetric, + of "M\ measure := \ A. \\<^isup>+x. f x * indicator A x \M \"]) + apply (simp_all add: eq) + apply (rule positive_integral_translated_density) + using f h by auto + then have "(\\<^isup>+x. f x * h x \M) \ \" using h(2) by simp then show "(\x. f x * h x) -` {\} \ space M \ null_sets" using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage) @@ -981,7 +1000,9 @@ next assume AE: "AE x. f x \ \" from sigma_finite guess Q .. note Q = this - interpret \: measure_space M \ by fact + interpret \: measure_space ?N + where "borel_measurable ?N = borel_measurable M" and "measure ?N = \" + and "sets ?N = sets M" and "space ?N = space M" using \ by (auto simp: measurable_def) def A \ "\i. f -` (case i of 0 \ {\} | Suc n \ {.. of_nat (Suc n)}) \ space M" { fix i j have "A i \ Q j \ sets M" unfolding A_def using f Q @@ -989,11 +1010,11 @@ by (cases i) (auto intro: measurable_sets[OF f]) } note A_in_sets = this let "?A n" = "case prod_decode n of (i,j) \ A i \ Q j" - show "sigma_finite_measure M \" + show "sigma_finite_measure ?N" proof (default, intro exI conjI subsetI allI) fix x assume "x \ range ?A" then obtain n where n: "x = ?A n" by auto - then show "x \ sets M" using A_in_sets by (cases "prod_decode n") auto + then show "x \ sets ?N" using A_in_sets by (cases "prod_decode n") auto next have "(\i. ?A i) = (\i j. A i \ Q j)" proof safe @@ -1014,16 +1035,16 @@ then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"]) qed qed (auto simp: A_def) - finally show "(\i. ?A i) = space M" by simp + finally show "(\i. ?A i) = space ?N" by simp next fix n obtain i j where [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto - have "(\\<^isup>+x. f x * indicator (A i \ Q j) x) \ \" + have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) \ \" proof (cases i) case 0 have "AE x. f x * indicator (A i \ Q j) x = 0" using AE by (rule AE_mp) (auto intro!: AE_cong simp: A_def `i = 0`) - then have "(\\<^isup>+x. f x * indicator (A i \ Q j) x) = 0" + then have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) = 0" using A_in_sets f apply (subst positive_integral_0_iff) apply fast @@ -1034,8 +1055,8 @@ then show ?thesis by simp next case (Suc n) - then have "(\\<^isup>+x. f x * indicator (A i \ Q j) x) \ - (\\<^isup>+x. of_nat (Suc n) * indicator (Q j) x)" + then have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) \ + (\\<^isup>+x. of_nat (Suc n) * indicator (Q j) x \M)" by (auto intro!: positive_integral_mono simp: indicator_def A_def) also have "\ = of_nat (Suc n) * \ (Q j)" using Q by (auto intro!: positive_integral_cmult_indicator) @@ -1043,33 +1064,34 @@ using Q by auto finally show ?thesis by simp qed - then show "\ (?A n) \ \" + then show "measure ?N (?A n) \ \" using A_in_sets Q eq by auto qed qed section "Radon-Nikodym derivative" -definition (in sigma_finite_measure) - "RN_deriv \ \ SOME f. f \ borel_measurable M \ - (\A \ sets M. \ A = (\\<^isup>+x. f x * indicator A x))" +definition + "RN_deriv M \ \ SOME f. f \ borel_measurable M \ + (\A \ sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" lemma (in sigma_finite_measure) RN_deriv_cong: - assumes cong: "\A. A \ sets M \ \' A = \ A" "\A. A \ sets M \ \' A = \ A" - shows "sigma_finite_measure.RN_deriv M \' \' x = RN_deriv \ x" + assumes cong: "\A. A \ sets M \ measure M' A = \ A" "sets M' = sets M" "space M' = space M" + and \: "\A. A \ sets M \ \' A = \ A" + shows "RN_deriv M' \' x = RN_deriv M \ x" proof - - interpret \': sigma_finite_measure M \' - using cong(1) by (rule sigma_finite_measure_cong) + interpret \': sigma_finite_measure M' + using cong by (rule sigma_finite_measure_cong) show ?thesis - unfolding RN_deriv_def \'.RN_deriv_def - by (simp add: cong positive_integral_cong_measure[OF cong(1)]) + unfolding RN_deriv_def + by (simp add: cong \ positive_integral_cong_measure[OF cong] measurable_def) qed lemma (in sigma_finite_measure) RN_deriv: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" assumes "absolutely_continuous \" - shows "RN_deriv \ \ borel_measurable M" (is ?borel) - and "\A. A \ sets M \ \ A = (\\<^isup>+x. RN_deriv \ x * indicator A x)" + shows "RN_deriv M \ \ borel_measurable M" (is ?borel) + and "\A. A \ sets M \ \ A = (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)" (is "\A. _ \ ?int A") proof - note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] @@ -1080,87 +1102,92 @@ qed lemma (in sigma_finite_measure) RN_deriv_positive_integral: - assumes \: "measure_space M \" "absolutely_continuous \" + assumes \: "measure_space (M\measure := \\)" "absolutely_continuous \" and f: "f \ borel_measurable M" - shows "measure_space.positive_integral M \ f = (\\<^isup>+x. RN_deriv \ x * f x)" + shows "integral\<^isup>P (M\measure := \\) f = (\\<^isup>+x. RN_deriv M \ x * f x \M)" proof - - interpret \: measure_space M \ by fact - have "\.positive_integral f = - measure_space.positive_integral M (\A. (\\<^isup>+x. RN_deriv \ x * indicator A x)) f" - by (intro \.positive_integral_cong_measure[symmetric] RN_deriv(2)[OF \, symmetric]) - also have "\ = (\\<^isup>+x. RN_deriv \ x * f x)" - by (intro positive_integral_translated_density RN_deriv[OF \] f) + interpret \: measure_space "M\measure := \\" by fact + have "integral\<^isup>P (M\measure := \\) f = + integral\<^isup>P (M\measure := \A. (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)\) f" + by (intro \.positive_integral_cong_measure[symmetric]) + (simp_all add: RN_deriv(2)[OF \, symmetric]) + also have "\ = (\\<^isup>+x. RN_deriv M \ x * f x \M)" + by (intro positive_integral_translated_density) + (simp_all add: RN_deriv[OF \] f) finally show ?thesis . qed lemma (in sigma_finite_measure) RN_deriv_unique: - assumes \: "measure_space M \" "absolutely_continuous \" + assumes \: "measure_space (M\measure := \\)" "absolutely_continuous \" and f: "f \ borel_measurable M" - and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x)" - shows "AE x. f x = RN_deriv \ x" + and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x \M)" + shows "AE x. f x = RN_deriv M \ x" proof (rule density_unique[OF f RN_deriv(1)[OF \]]) fix A assume A: "A \ sets M" - show "(\\<^isup>+x. f x * indicator A x) = (\\<^isup>+x. RN_deriv \ x * indicator A x)" + show "(\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)" unfolding eq[OF A, symmetric] RN_deriv(2)[OF \ A, symmetric] .. qed - lemma (in sigma_finite_measure) RN_deriv_finite: - assumes sfm: "sigma_finite_measure M \" and ac: "absolutely_continuous \" - shows "AE x. RN_deriv \ x \ \" + assumes sfm: "sigma_finite_measure (M\measure := \\)" and ac: "absolutely_continuous \" + shows "AE x. RN_deriv M \ x \ \" proof - - interpret \: sigma_finite_measure M \ by fact - have \: "measure_space M \" by default + interpret \: sigma_finite_measure "M\measure := \\" by fact + have \: "measure_space (M\measure := \\)" by default from sfm show ?thesis using sigma_finite_iff_density_finite[OF \ RN_deriv[OF \ ac]] by simp qed lemma (in sigma_finite_measure) - assumes \: "sigma_finite_measure M \" "absolutely_continuous \" + assumes \: "sigma_finite_measure (M\measure := \\)" "absolutely_continuous \" and f: "f \ borel_measurable M" - shows RN_deriv_integral: "measure_space.integral M \ f = (\x. real (RN_deriv \ x) * f x)" (is ?integral) - and RN_deriv_integrable: "measure_space.integrable M \ f \ integrable (\x. real (RN_deriv \ x) * f x)" (is ?integrable) + shows RN_deriv_integrable: "integrable (M\measure := \\) f \ + integrable M (\x. real (RN_deriv M \ x) * f x)" (is ?integrable) + and RN_deriv_integral: "integral\<^isup>L (M\measure := \\) f = + (\x. real (RN_deriv M \ x) * f x \M)" (is ?integral) proof - - interpret \: sigma_finite_measure M \ by fact - have ms: "measure_space M \" by default + interpret \: sigma_finite_measure "M\measure := \\" by fact + have ms: "measure_space (M\measure := \\)" by default have minus_cong: "\A B A' B'::pextreal. A = A' \ B = B' \ real A - real B = real A' - real B'" by simp have f': "(\x. - f x) \ borel_measurable M" using f by auto - { fix f :: "'a \ real" assume "f \ borel_measurable M" - { fix x assume *: "RN_deriv \ x \ \" - have "Real (real (RN_deriv \ x)) * Real (f x) = Real (real (RN_deriv \ x) * f x)" + have Nf: "f \ borel_measurable (M\measure := \\)" using f unfolding measurable_def by auto + { fix f :: "'a \ real" + { fix x assume *: "RN_deriv M \ x \ \" + have "Real (real (RN_deriv M \ x)) * Real (f x) = Real (real (RN_deriv M \ x) * f x)" by (simp add: mult_le_0_iff) - then have "RN_deriv \ x * Real (f x) = Real (real (RN_deriv \ x) * f x)" + then have "RN_deriv M \ x * Real (f x) = Real (real (RN_deriv M \ x) * f x)" using * by (simp add: Real_real) } note * = this - have "(\\<^isup>+x. RN_deriv \ x * Real (f x)) = (\\<^isup>+x. Real (real (RN_deriv \ x) * f x))" + have "(\\<^isup>+x. RN_deriv M \ x * Real (f x) \M) = (\\<^isup>+x. Real (real (RN_deriv M \ x) * f x) \M)" apply (rule positive_integral_cong_AE) apply (rule AE_mp[OF RN_deriv_finite[OF \]]) by (auto intro!: AE_cong simp: *) } - with this[OF f] this[OF f'] f f' + with this this f f' Nf show ?integral ?integrable - unfolding \.integral_def integral_def \.integrable_def integrable_def - by (auto intro!: RN_deriv(1)[OF ms \(2)] minus_cong simp: RN_deriv_positive_integral[OF ms \(2)]) + unfolding lebesgue_integral_def integrable_def + by (auto intro!: RN_deriv(1)[OF ms \(2)] minus_cong + simp: RN_deriv_positive_integral[OF ms \(2)]) qed lemma (in sigma_finite_measure) RN_deriv_singleton: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" and ac: "absolutely_continuous \" and "{x} \ sets M" - shows "\ {x} = RN_deriv \ x * \ {x}" + shows "\ {x} = RN_deriv M \ x * \ {x}" proof - note deriv = RN_deriv[OF assms(1, 2)] from deriv(2)[OF `{x} \ sets M`] - have "\ {x} = (\\<^isup>+w. RN_deriv \ x * indicator {x} w)" + have "\ {x} = (\\<^isup>+w. RN_deriv M \ x * indicator {x} w \M)" by (auto simp: indicator_def intro!: positive_integral_cong) thus ?thesis using positive_integral_cmult_indicator[OF `{x} \ sets M`] by auto qed theorem (in finite_measure_space) RN_deriv_finite_measure: - assumes "measure_space M \" + assumes "measure_space (M\measure := \\)" and ac: "absolutely_continuous \" and "x \ space M" - shows "\ {x} = RN_deriv \ x * \ {x}" + shows "\ {x} = RN_deriv M \ x * \ {x}" proof - have "{x} \ sets M" using sets_eq_Pow `x \ space M` by auto from RN_deriv_singleton[OF assms(1,2) this] show ?thesis . diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Wed Feb 02 12:34:45 2011 +0100 @@ -29,7 +29,7 @@ sets :: "'a set set" locale algebra = - fixes M :: "'a algebra" + fixes M :: "('a, 'b) algebra_scheme" assumes space_closed: "sets M \ Pow (space M)" and empty_sets [iff]: "{} \ sets M" and compl_sets [intro]: "!!a. a \ sets M \ space M - a \ sets M" @@ -104,7 +104,7 @@ section {* Restricted algebras *} abbreviation (in algebra) - "restricted_space A \ \ space = A, sets = (\S. (A \ S)) ` sets M \" + "restricted_space A \ \ space = A, sets = (\S. (A \ S)) ` sets M, \ = more M \" lemma (in algebra) restricted_algebra: assumes "A \ sets M" shows "algebra (restricted_space A)" @@ -222,7 +222,7 @@ | Union: "(\i::nat. a i \ sigma_sets sp A) \ (\i. a i) \ sigma_sets sp A" definition - "sigma M = \ space = space M, sets = sigma_sets (space M) (sets M) \" + "sigma M = \ space = space M, sets = sigma_sets (space M) (sets M), \ = more M \" lemma (in sigma_algebra) sigma_sets_subset: assumes a: "a \ sets M" @@ -354,12 +354,12 @@ qed lemma sigma_sets_Int: - assumes "A \ sigma_sets sp st" - shows "op \ A ` sigma_sets sp st = sigma_sets (A \ sp) (op \ A ` st)" + assumes "A \ sigma_sets sp st" "A \ sp" + shows "op \ A ` sigma_sets sp st = sigma_sets A (op \ A ` st)" proof (intro equalityI subsetI) fix x assume "x \ op \ A ` sigma_sets sp st" then obtain y where "y \ sigma_sets sp st" "x = y \ A" by auto - then show "x \ sigma_sets (A \ sp) (op \ A ` st)" + then have "x \ sigma_sets (A \ sp) (op \ A ` st)" proof (induct arbitrary: x) case (Compl a) then show ?case @@ -370,13 +370,15 @@ by (auto intro!: sigma_sets.Union simp add: UN_extend_simps simp del: UN_simps) qed (auto intro!: sigma_sets.intros) + then show "x \ sigma_sets A (op \ A ` st)" + using `A \ sp` by (simp add: Int_absorb2) next - fix x assume "x \ sigma_sets (A \ sp) (op \ A ` st)" + fix x assume "x \ sigma_sets A (op \ A ` st)" then show "x \ op \ A ` sigma_sets sp st" proof induct case (Compl a) then obtain x where "a = A \ x" "x \ sigma_sets sp st" by auto - then show ?case + then show ?case using `A \ sp` by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl) next case (Union a) @@ -707,7 +709,7 @@ subsection {* Sigma algebra generated by function preimages *} definition (in sigma_algebra) - "vimage_algebra S f = \ space = S, sets = (\A. f -` A \ S) ` sets M \" + "vimage_algebra S f = \ space = S, sets = (\A. f -` A \ S) ` sets M, \ = more M \" lemma (in sigma_algebra) in_vimage_algebra[simp]: "A \ sets (vimage_algebra S f) \ (\B\sets M. A = f -` B \ S)" @@ -870,7 +872,7 @@ section {* Conditional space *} definition (in algebra) - "image_space X = \ space = X`space M, sets = (\S. X`S) ` sets M \" + "image_space X = \ space = X`space M, sets = (\S. X`S) ` sets M, \ = more M \" definition (in algebra) "conditional_space X A = algebra.image_space (restricted_space A) X" @@ -1158,7 +1160,7 @@ section {* Dynkin systems *} locale dynkin_system = - fixes M :: "'a algebra" + fixes M :: "('a, 'b) algebra_scheme" assumes space_closed: "sets M \ Pow (space M)" and space: "space M \ sets M" and compl[intro!]: "\A. A \ sets M \ space M - A \ sets M" @@ -1239,12 +1241,12 @@ subsection "Smallest Dynkin systems" -definition dynkin :: "'a algebra \ 'a algebra" where +definition dynkin where "dynkin M = \ space = space M, - sets = \{D. dynkin_system \ space = space M, sets = D\ \ sets M \ D}\" + sets = \{D. dynkin_system \ space = space M, sets = D \ \ sets M \ D}, + \ = more M \" lemma dynkin_system_dynkin: - fixes M :: "'a algebra" assumes "sets M \ Pow (space M)" shows "dynkin_system (dynkin M)" proof (rule dynkin_systemI) @@ -1311,21 +1313,17 @@ qed lemma (in dynkin_system) dynkin_subset: - fixes N :: "'a algebra" assumes "sets N \ sets M" assumes "space N = space M" shows "sets (dynkin N) \ sets M" proof - - have *: "\space = space N, sets = sets M\ = M" - unfolding `space N = space M` by simp have "dynkin_system M" by default - then have "dynkin_system \space = space N, sets = sets M\" - using assms unfolding * by simp + then have "dynkin_system \space = space N, sets = sets M \" + using assms unfolding dynkin_system_def by simp with `sets N \ sets M` show ?thesis by (auto simp add: dynkin_def) qed lemma sigma_eq_dynkin: - fixes M :: "'a algebra" assumes sets: "sets M \ Pow (space M)" assumes "Int_stable M" shows "sigma M = dynkin M" @@ -1367,7 +1365,7 @@ have "sigma_sets (space M) (sets M) \ sets (dynkin M)" by auto ultimately have "sigma_sets (space M) (sets M) = sets (dynkin M)" by auto then show ?thesis - by (intro algebra.equality) (simp_all add: sigma_def) + by (auto intro!: algebra.equality simp: sigma_def dynkin_def) qed lemma (in dynkin_system) dynkin_idem: @@ -1381,23 +1379,22 @@ by (intro dynkin_subset) auto qed then show ?thesis - by (auto intro!: algebra.equality) + by (auto intro!: algebra.equality simp: dynkin_def) qed lemma (in dynkin_system) dynkin_lemma: - fixes E :: "'a algebra" - assumes "Int_stable E" and E: "sets E \ sets M" "space E = space M" - and "sets M \ sets (sigma E)" - shows "sigma E = M" + assumes "Int_stable E" + and E: "sets E \ sets M" "space E = space M" "sets M \ sets (sigma E)" + shows "sets (sigma E) = sets M" proof - have "sets E \ Pow (space E)" - using E sets_into_space by auto + using E sets_into_space by force then have "sigma E = dynkin E" using `Int_stable E` by (rule sigma_eq_dynkin) moreover then have "sets (dynkin E) = sets M" - using assms dynkin_subset[OF E] by simp + using assms dynkin_subset[OF E(1,2)] by simp ultimately show ?thesis - using E by simp + using assms by (auto intro!: algebra.equality simp: dynkin_def) qed subsection "Sigma algebras on finite sets" @@ -1467,6 +1464,7 @@ lemma vimage_algebra_sigma: assumes bi: "bij_inv (space (sigma F)) (space (sigma E)) f e" and "sets E \ Pow (space E)" and F: "sets F \ Pow (space F)" + and "more E = more F" and "f \ measurable F E" "e \ measurable E F" shows "sigma_algebra.vimage_algebra (sigma E) (space (sigma F)) f = sigma F" proof - @@ -1486,8 +1484,8 @@ using `f \ measurable F E` unfolding measurable_def by auto qed show "vimage_algebra (space (sigma F)) f = sigma F" - unfolding vimage_algebra_def - using assms by (auto simp: bij_inv_def eq sigma_sets_vimage[symmetric] sigma_def) + unfolding vimage_algebra_def using assms + by (auto simp: bij_inv_def eq sigma_sets_vimage[symmetric] sigma_def) qed lemma measurable_sigma_sigma: diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Wed Feb 02 12:34:45 2011 +0100 @@ -7,25 +7,25 @@ assumes finite[simp]: "finite S" and not_empty[simp]: "S \ {}" -definition (in finite_space) "M = \ space = S, sets = Pow S \" -definition (in finite_space) \_def[simp]: "\ A = (of_nat (card A) / of_nat (card S) :: pextreal)" +definition (in finite_space) "M = \ space = S, sets = Pow S, + measure = (\A. of_nat (card A) / of_nat (card S) :: pextreal) \" lemma (in finite_space) shows space_M[simp]: "space M = S" and sets_M[simp]: "sets M = Pow S" by (simp_all add: M_def) -sublocale finite_space \ finite_measure_space M \ +sublocale finite_space \ finite_measure_space M proof (rule finite_measure_spaceI) fix A B :: "'a set" assume "A \ B = {}" "A \ space M" "B \ space M" - then show "\ (A \ B) = \ A + \ B" - by (simp add: inverse_eq_divide field_simps Real_real + then show "measure M (A \ B) = measure M A + measure M B" + by (simp add: inverse_eq_divide field_simps Real_real M_def divide_le_0_iff zero_le_divide_iff card_Un_disjoint finite_subset[OF _ finite]) -qed auto +qed (auto simp: M_def) -sublocale finite_space \ information_space M \ 2 - by default simp_all +sublocale finite_space \ information_space M 2 + by default (simp_all add: M_def) lemma set_of_list_extend: "{xs. length xs = Suc n \ (\x\set xs. x \ A)} = @@ -506,7 +506,8 @@ moreover from x have "length x = n" by (auto simp: inversion_def_raw dc_crypto) ultimately have "real (?dIP {(x, z)}) = 2 / (real n * 2^n)" using x - by (simp add: distribution_def card_dc_crypto card_payer_and_inversion + unfolding distribution_def + by (simp add: M_def card_dc_crypto card_payer_and_inversion inverse_eq_divide mult_le_0_iff zero_le_mult_iff power_le_zero_eq) moreover from z have "payer -` {z} \ dc_crypto = {z} \ {xs. length xs = n}" @@ -514,8 +515,8 @@ hence "card (payer -` {z} \ dc_crypto) = 2^n" using card_list_length[where A="UNIV::bool set"] by (simp add: card_cartesian_product_singleton) - hence "real (?dP {z}) = 1 / real n" - by (simp add: distribution_def card_dc_crypto field_simps inverse_eq_divide + hence "real (?dP {z}) = 1 / real n" unfolding distribution_def + by (simp add: card_dc_crypto field_simps inverse_eq_divide M_def mult_le_0_iff zero_le_mult_iff power_le_zero_eq) ultimately show "real (?dIP {(x,z)}) * log 2 (real (?dIP {(x,z)}) / real (?dP {z})) = @@ -535,8 +536,8 @@ fix x assume x_inv: "x \ inversion ` dc_crypto" hence "length x = n" by (auto simp: inversion_def_raw dc_crypto) moreover have "inversion -` {x} \ dc_crypto = {dc \ dc_crypto. inversion dc = x}" by auto - ultimately have "?dI {x} = 2 / 2^n" using `0 < n` - by (simp add: distribution_def card_inversion[OF x_inv] card_dc_crypto + ultimately have "?dI {x} = 2 / 2^n" using `0 < n` unfolding distribution_def + by (simp add: card_inversion[OF x_inv] card_dc_crypto M_def mult_le_0_iff zero_le_mult_iff power_le_zero_eq) thus "real (?dI {x}) * log 2 (real (?dI {x})) = 2 * (1 - real n) / 2^n" by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide) diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Feb 02 12:34:45 2011 +0100 @@ -1,6 +1,6 @@ -(* Author: Johannes Hoelzl, TU Muenchen *) +(* Author: Johannes Hölzl, TU München *) -header {* Formalization of a Countermeasure by Koepf & Duermuth 2009 *} +header {* Formalization of a Countermeasure by Koepf \& Duermuth 2009 *} theory Koepf_Duermuth_Countermeasure imports Information "~~/src/HOL/Library/Permutation" @@ -80,9 +80,7 @@ shows "bij_betw (ordered_variable_partition X) {0..i j. \ i < card (X`space M) ; j < card (X`space M) ; i \ j \ \ distribution X {ordered_variable_partition X i} \ distribution X {ordered_variable_partition X j}" -proof - - -qed + oops definition (in prob_space) "order_distribution X i = real (distribution X {ordered_variable_partition X i})" @@ -90,12 +88,10 @@ definition (in prob_space) "guessing_entropy b X = (\i'(_')") where "\(X) \ guessing_entropy b X" - - lemma zero_notin_Suc_image[simp]: "0 \ Suc ` A" by auto @@ -103,7 +99,7 @@ "extensional d A = {f. \x. x \ A \ f x = d}" lemma extensional_empty[simp]: "extensional d {} = {\x. d}" - unfolding extensional_def by (simp add: expand_set_eq expand_fun_eq) + unfolding extensional_def by (simp add: set_eq_iff fun_eq_iff) lemma funset_eq_UN_fun_upd_I: assumes *: "\f. f \ F (insert a A) \ f(a := d) \ F A" @@ -138,7 +134,7 @@ using assms by induct auto lemma fun_upd_eq_iff: "f(a := b) = g(a := b') \ b = b' \ f(a := d) = g(a := d)" - by (auto simp: expand_fun_eq) + by (auto simp: fun_eq_iff) lemma card_funcset: assumes "finite A" "finite B" @@ -205,24 +201,14 @@ lemma (in finite_information) positive_p_sum[simp]: "0 \ setsum p X" by (auto intro!: setsum_nonneg) -sublocale finite_information \ finite_information_space "\ space = \, sets = Pow \ \" "\S. Real (setsum p S)" b -proof - - show "finite_information_space \ space = \, sets = Pow \ \ (\S. Real (setsum p S)) b" - unfolding finite_information_space_def finite_information_space_axioms_def - unfolding finite_prob_space_def prob_space_def prob_space_axioms_def - unfolding finite_measure_space_def finite_measure_space_axioms_def - by (force intro!: sigma_algebra.finite_additivity_sufficient - simp: additive_def sigma_algebra_Pow positive_def Real_eq_Real - setsum.union_disjoint finite_subset) -qed +sublocale finite_information \ finite_measure_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" + by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset) -lemma (in prob_space) prob_space_subalgebra: - assumes "N \ sets M" "sigma_algebra (M\ sets := N \)" - shows "prob_space (M\ sets := N \) \" sorry +sublocale finite_information \ finite_prob_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" + by default simp -lemma (in measure_space) measure_space_subalgebra: - assumes "N \ sets M" "sigma_algebra (M\ sets := N \)" - shows "measure_space (M\ sets := N \) \" sorry +sublocale finite_information \ information_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" b + by default simp locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b for b :: real @@ -258,12 +244,11 @@ next case (Suc n) then show ?case - by (simp add: comp_def set_of_list_extend - lessThan_eq_Suc_image setsum_reindex setprod_reindex) + by (simp add: comp_def set_of_list_extend lessThan_Suc_eq_insert_0 + setsum_reindex setprod_reindex) qed then show "setsum P msgs = 1" unfolding msgs_def P_def by simp - fix x have "\ A f. 0 \ (\x\A. M (f x))" by (auto simp: setprod_nonneg) then show "0 \ P x" @@ -303,7 +288,7 @@ show ?case unfolding * using Suc[of "\i. P (Suc i)"] by (simp add: setsum_reindex split_conv setsum_cartesian_product' - lessThan_eq_Suc_image setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric]) + lessThan_Suc_eq_insert_0 setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric]) qed context koepf_duermuth @@ -347,22 +332,22 @@ "\

(X|Y) x \ \

(X, Y) x / \

(Y) (snd`x)" notation - finite_entropy ("\'( _ ')") + entropy_Pow ("\'( _ ')") notation - finite_conditional_entropy ("\'( _ | _ ')") + conditional_entropy_Pow ("\'( _ | _ ')") notation - finite_mutual_information ("\'( _ ; _ ')") + mutual_information_Pow ("\'( _ ; _ ')") lemma t_eq_imp_bij_func: assumes "t xs = t ys" shows "\f. bij_betw f {.. (\i ys" unfolding multiset_of_eq_perm count_inject . - then show ?thesis by (rule permutation_Ex_func) + then show ?thesis by (rule permutation_Ex_bij) qed lemma \

_k: assumes "k \ keys" shows "\

(fst) {k} = K k" @@ -471,8 +456,6 @@ let "?H obs" = "(\k\keys. \

(fst|OB) {(k, obs)} * log b (\

(fst|OB) {(k, obs)})) :: real" let "?Ht obs" = "(\k\keys. \

(fst|t\OB) {(k, obs)} * log b (\

(fst|t\OB) {(k, obs)})) :: real" - note [[simproc del: finite_information_space.mult_log]] - { fix obs assume obs: "obs \ OB`msgs" have "?H obs = (\k\keys. \

(fst|t\OB) {(k, t obs)} * log b (\

(fst|t\OB) {(k, t obs)}))" using CP_T_eq_CP_O[OF _ obs] @@ -495,7 +478,8 @@ txt {* Lemma 3 *} have "\(fst | OB) = -(\obs\OB`msgs. \

(OB) {obs} * ?Ht (t obs))" - unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp + unfolding conditional_entropy_eq_ce_with_hypothesis[OF + simple_function_finite simple_function_finite] using * by simp also have "\ = -(\obs\t`OB`msgs. \

(t\OB) {obs} * ?Ht obs)" apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t]) apply (subst setsum_reindex) @@ -509,20 +493,22 @@ by (simp add: setsum_divide_distrib[symmetric] field_simps ** setsum_right_distrib[symmetric] setsum_left_distrib[symmetric]) also have "\ = \(fst | t\OB)" - unfolding conditional_entropy_eq_ce_with_hypothesis + unfolding conditional_entropy_eq_ce_with_hypothesis[OF + simple_function_finite simple_function_finite] by (simp add: comp_def image_image[symmetric]) finally show ?thesis . qed theorem "\(fst ; OB) \ real (card observations) * log b (real n + 1)" proof - + from simple_function_finite simple_function_finite have "\(fst ; OB) = \(fst) - \(fst | OB)" - using mutual_information_eq_entropy_conditional_entropy . + by (rule mutual_information_eq_entropy_conditional_entropy) also have "\ = \(fst) - \(fst | t\OB)" unfolding ce_OB_eq_ce_t .. also have "\ = \(t\OB) - \(t\OB | fst)" - unfolding entropy_chain_rule[symmetric] sign_simps - by (subst entropy_commute) simp + unfolding entropy_chain_rule[symmetric, OF simple_function_finite simple_function_finite] sign_simps + by (subst entropy_commute[OF simple_function_finite simple_function_finite]) simp also have "\ \ \(t\OB)" using conditional_entropy_positive[of "t\OB" fst] by simp also have "\ \ log b (real (card ((t\OB)`msgs)))" @@ -537,4 +523,4 @@ end -end \ No newline at end of file +end