# HG changeset patch # User hoelzl # Date 1335176075 -7200 # Node ID 05663f75964c70583a5bb479abc20ae42eab12c0 # Parent 64023cf4d1488accdb04d15acb677ee6ce45d7a9 reworked Probability theory diff -r 64023cf4d148 -r 05663f75964c NEWS --- a/NEWS Mon Apr 23 12:23:23 2012 +0100 +++ b/NEWS Mon Apr 23 12:14:35 2012 +0200 @@ -647,6 +647,180 @@ DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing +* Theory Library/Multiset: Improved code generation of multisets. + +* Session HOL-Probability: Introduced the type "'a measure" to represent +measures, this replaces the records 'a algebra and 'a measure_space. The +locales based on subset_class now have two locale-parameters the space +\ and the set of measurables sets M. The product of probability spaces +uses now the same constant as the finite product of sigma-finite measure +spaces "PiM :: ('i => 'a) measure". Most constants are defined now +outside of locales and gain an additional parameter, like null_sets, +almost_eventually or \'. Measure space constructions for distributions +and densities now got their own constants distr and density. Instead of +using locales to describe measure spaces with a finite space, the +measure count_space and point_measure is introduced. INCOMPATIBILITY. + + Renamed constants: + measure -> emeasure + finite_measure.\' -> measure + product_algebra_generator -> prod_algebra + product_prob_space.emb -> prod_emb + product_prob_space.infprod_algebra -> PiM + + Removed locales: + completeable_measure_space + finite_measure_space + finite_prob_space + finite_product_finite_prob_space + finite_product_sigma_algebra + finite_sigma_algebra + measure_space + pair_finite_prob_space + pair_finite_sigma_algebra + pair_finite_space + pair_sigma_algebra + product_sigma_algebra + + Removed constants: + distribution -> use distr measure, or distributed predicate + joint_distribution -> use distr measure, or distributed predicate + product_prob_space.infprod_algebra -> use PiM + subvimage + image_space + conditional_space + pair_measure_generator + + Replacement theorems: + sigma_algebra.measurable_sigma -> measurable_measure_of + measure_space.additive -> emeasure_additive + measure_space.measure_additive -> plus_emeasure + measure_space.measure_mono -> emeasure_mono + measure_space.measure_top -> emeasure_space + measure_space.measure_compl -> emeasure_compl + measure_space.measure_Diff -> emeasure_Diff + measure_space.measure_countable_increasing -> emeasure_countable_increasing + measure_space.continuity_from_below -> SUP_emeasure_incseq + measure_space.measure_incseq -> incseq_emeasure + measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq + measure_space.measure_decseq -> decseq_emeasure + measure_space.continuity_from_above -> INF_emeasure_decseq + measure_space.measure_insert -> emeasure_insert + measure_space.measure_setsum -> setsum_emeasure + measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton + finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite + measure_space.measure_setsum_split -> setsum_emeasure_cover + measure_space.measure_subadditive -> subadditive + measure_space.measure_subadditive_finite -> emeasure_subadditive_finite + measure_space.measure_eq_0 -> emeasure_eq_0 + measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite + measure_space.measure_countably_subadditive -> emeasure_subadditive_countably + measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0 + measure_unique_Int_stable -> measure_eqI_generator_eq + measure_space.measure_Diff_null_set -> emeasure_Diff_null_set + measure_space.measure_Un_null_set -> emeasure_Un_null_set + measure_space.almost_everywhere_def -> eventually_ae_filter + measure_space.almost_everywhere_vimage -> AE_distrD + measure_space.measure_space_vimage -> emeasure_distr + measure_space.AE_iff_null_set -> AE_iff_null + measure_space.real_measure_Union -> measure_Union + measure_space.real_measure_finite_Union -> measure_finite_Union + measure_space.real_measure_Diff -> measure_Diff + measure_space.real_measure_UNION -> measure_UNION + measure_space.real_measure_subadditive -> measure_subadditive + measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton + measure_space.real_continuity_from_below -> Lim_measure_incseq + measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq + measure_space.real_continuity_from_above -> Lim_measure_decseq + measure_space.real_measure_countably_subadditive -> measure_subadditive_countably + finite_measure.finite_measure -> finite_measure.emeasure_finite + finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure + finite_measure.positive_measure' -> measure_nonneg + finite_measure.real_measure -> finite_measure.emeasure_real + finite_measure.empty_measure -> measure_empty + finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably + finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton + finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq + finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq + measure_space.simple_integral_vimage -> simple_integral_distr + measure_space.integrable_vimage -> integrable_distr + measure_space.positive_integral_translated_density -> positive_integral_density + measure_space.integral_translated_density -> integral_density + measure_space.integral_vimage -> integral_distr + measure_space_density -> emeasure_density + measure_space.positive_integral_vimage -> positive_integral_distr + measure_space.simple_function_vimage -> simple_function_comp + measure_space.simple_integral_vimage -> simple_integral_distr + pair_sigma_algebra.measurable_cut_fst -> sets_Pair1 + pair_sigma_algebra.measurable_cut_snd -> sets_Pair2 + pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1 + pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2 + pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff + pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1 + pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2 + measure_space.measure_not_negative -> emeasure_not_MInf + pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap + pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt + pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2 + pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times + pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap + pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap' + pair_sigma_algebra.sets_swap -> sets_pair_swap + finite_product_sigma_algebra.in_P -> sets_PiM_I_finite + Int_stable_product_algebra_generator -> positive_integral + product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge + product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton + product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge + finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty + product_algebra_generator_der -> prod_algebra_eq_finite + product_algebra_generator_into_space -> prod_algebra_sets_into_space + product_sigma_algebra.product_algebra_into_space -> space_closed + product_algebraE -> prod_algebraE_all + product_algebraI -> sets_PiM_I_finite + product_measure_exists -> product_sigma_finite.sigma_finite + sets_product_algebra -> sets_PiM + sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq + space_product_algebra -> space_PiM + Int_stable_cuboids -> Int_stable_atLeastAtMost + measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density + sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr + prob_space_unique_Int_stable -> measure_eqI_prob_space + sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint + prob_space.measure_space_1 -> prob_space.emeasure_space_1 + prob_space.prob_space_vimage -> prob_space_distr + prob_space.random_variable_restrict -> measurable_restrict + measure_preserving -> equality "distr M N f = N" "f : measurable M N" + measure_unique_Int_stable_vimage -> measure_eqI_generator_eq + measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq + product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator + product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb + finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb + product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty + product_prob_space.measurable_component -> measurable_component_singleton + product_prob_space.measurable_emb -> measurable_prod_emb + product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single + product_prob_space.measurable_singleton_infprod -> measurable_component_singleton + product_prob_space.measure_emb -> emeasure_prod_emb + sequence_space.measure_infprod -> sequence_space.measure_PiM_countable + product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict + prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM + prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq + conditional_entropy_positive -> conditional_entropy_nonneg_simple + conditional_entropy_eq -> conditional_entropy_simple_distributed + conditional_mutual_information_eq_mutual_information -> conditional_mutual_information_eq_mutual_information_simple + conditional_mutual_information_generic_positive -> conditional_mutual_information_nonneg_simple + conditional_mutual_information_positive -> conditional_mutual_information_nonneg_simple + entropy_commute -> entropy_commute_simple + entropy_eq -> entropy_simple_distributed + entropy_generic_eq -> entropy_simple_distributed + entropy_positive -> entropy_nonneg_simple + entropy_uniform_max -> entropy_uniform + KL_eq_0 -> KL_same_eq_0 + KL_eq_0_imp -> KL_eq_0_iff_eq + KL_ge_0 -> KL_nonneg + mutual_information_eq -> mutual_information_simple_distributed + mutual_information_positive -> mutual_information_nonneg_simple + * New "case_product" attribute to generate a case rule doing multiple case distinctions at the same time. E.g. @@ -655,8 +829,6 @@ produces a rule which can be used to perform case distinction on both a list and a nat. -* Theory Library/Multiset: Improved code generation of multisets. - * New Transfer package: - transfer_rule attribute: Maintains a collection of transfer rules, diff -r 64023cf4d148 -r 05663f75964c src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/IsaMakefile Mon Apr 23 12:14:35 2012 +0200 @@ -1198,14 +1198,13 @@ $(OUT)/HOL-Probability: $(OUT)/HOL-Multivariate_Analysis \ Probability/Binary_Product_Measure.thy Probability/Borel_Space.thy \ Probability/Caratheodory.thy Probability/Complete_Measure.thy \ - Probability/Conditional_Probability.thy \ Probability/ex/Dining_Cryptographers.thy \ Probability/ex/Koepf_Duermuth_Countermeasure.thy \ Probability/Finite_Product_Measure.thy \ Probability/Independent_Family.thy \ Probability/Infinite_Product_Measure.thy Probability/Information.thy \ Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \ - Probability/Measure.thy Probability/Probability_Measure.thy \ + Probability/Measure_Space.thy Probability/Probability_Measure.thy \ Probability/Probability.thy Probability/Radon_Nikodym.thy \ Probability/ROOT.ML Probability/Sigma_Algebra.thy \ Library/Countable.thy Library/FuncSet.thy Library/Nat_Bijection.thy diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -28,413 +28,317 @@ section "Binary products" -definition - "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 :: "('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) - "P \ M1 \\<^isub>M M2" - -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.space_closed M2.space_closed - by (rule sigma_algebra_pair_measure) - -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 pair_measureI[intro, simp]: - "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" - by (auto simp add: pair_measure_def) + "A \\<^isub>M B = measure_of (space A \ space B) + {a \ b | a b. a \ sets A \ b \ sets B} + (\X. \\<^isup>+x. (\\<^isup>+y. indicator X (x,y) \B) \A)" lemma space_pair_measure: "space (A \\<^isub>M B) = space A \ space B" - by (simp add: pair_measure_def pair_measure_generator_def) + unfolding pair_measure_def using space_closed[of A] space_closed[of B] + by (intro space_measure_of) auto + +lemma sets_pair_measure: + "sets (A \\<^isub>M B) = sigma_sets (space A \ space B) {a \ b | a b. a \ sets A \ b \ sets B}" + unfolding pair_measure_def using space_closed[of A] space_closed[of B] + by (intro sets_measure_of) auto -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_measureI[intro, simp]: + "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" + by (auto simp: sets_pair_measure) -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 measurable_pair_measureI: + assumes 1: "f \ space M \ space M1 \ space M2" + assumes 2: "\A B. A \ sets M1 \ B \ sets M2 \ f -` (A \ B) \ space M \ sets M" + shows "f \ measurable M (M1 \\<^isub>M M2)" + unfolding pair_measure_def using 1 2 + by (intro measurable_measure_of) (auto dest: sets_into_space) -lemma pair_measure_generator_Int_snd: - assumes "sets S1 \ Pow (space S1)" - 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") - 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 measurable_fst[intro!, simp]: "fst \ measurable (M1 \\<^isub>M M2) M1" + unfolding measurable_def +proof safe + fix A assume A: "A \ sets M1" + from this[THEN sets_into_space] have "fst -` A \ space M1 \ space M2 = A \ space M2" by auto + with A show "fst -` A \ space (M1 \\<^isub>M M2) \ sets (M1 \\<^isub>M M2)" by (simp add: space_pair_measure) +qed (simp add: space_pair_measure) -lemma (in pair_sigma_algebra) - shows measurable_fst[intro!, simp]: - "fst \ measurable P M1" (is ?fst) - and measurable_snd[intro!, simp]: - "snd \ measurable P M2" (is ?snd) -proof - - { fix X assume "X \ sets M1" - then have "\X1\sets M1. \X2\sets M2. fst -` X \ space M1 \ space M2 = X1 \ X2" - apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) - using M1.sets_into_space by force+ } - moreover - { fix X assume "X \ sets M2" - then have "\X1\sets M1. \X2\sets M2. snd -` X \ space M1 \ space M2 = X1 \ X2" - apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) - using M2.sets_into_space by force+ } - ultimately have "?fst \ ?snd" - by (fastforce simp: measurable_def sets_sigma space_pair_measure - intro!: sigma_sets.Basic) - then show ?fst ?snd by auto -qed +lemma measurable_snd[intro!, simp]: "snd \ measurable (M1 \\<^isub>M M2) M2" + unfolding measurable_def +proof safe + fix A assume A: "A \ sets M2" + from this[THEN sets_into_space] have "snd -` A \ space M1 \ space M2 = space M1 \ A" by auto + with A show "snd -` A \ space (M1 \\<^isub>M M2) \ sets (M1 \\<^isub>M M2)" by (simp add: space_pair_measure) +qed (simp add: space_pair_measure) + +lemma measurable_fst': "f \ measurable M (N \\<^isub>M P) \ (\x. fst (f x)) \ measurable M N" + using measurable_comp[OF _ measurable_fst] by (auto simp: comp_def) + +lemma measurable_snd': "f \ measurable M (N \\<^isub>M P) \ (\x. snd (f x)) \ measurable M P" + using measurable_comp[OF _ measurable_snd] by (auto simp: comp_def) -lemma (in pair_sigma_algebra) measurable_pair_iff: - assumes "sigma_algebra M" - shows "f \ measurable M P \ - (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" -proof - - interpret M: sigma_algebra M by fact - 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" unfolding pair_measure_def - proof (rule M.measurable_sigma) - 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 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_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" - using s `C \ sets M2` unfolding measurable_def by auto - moreover have "f -` A \ space M = ((fst \ f) -` B \ space M) \ ((snd \ f) -` C \ space M)" - unfolding `A = B \ C` by (auto simp: vimage_Times) - ultimately show "f -` A \ space M \ sets M" by auto - qed +lemma measurable_fst'': "f \ measurable M N \ (\x. f (fst x)) \ measurable (M \\<^isub>M P) N" + using measurable_comp[OF measurable_fst _] by (auto simp: comp_def) + +lemma measurable_snd'': "f \ measurable M N \ (\x. f (snd x)) \ measurable (P \\<^isub>M M) N" + using measurable_comp[OF measurable_snd _] by (auto simp: comp_def) + +lemma measurable_pair_iff: + "f \ measurable M (M1 \\<^isub>M M2) \ (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" +proof safe + assume f: "(fst \ f) \ measurable M M1" and s: "(snd \ f) \ measurable M M2" + show "f \ measurable M (M1 \\<^isub>M M2)" + proof (rule measurable_pair_measureI) + show "f \ space M \ space M1 \ space M2" + using f s by (auto simp: mem_Times_iff measurable_def comp_def) + fix A B assume "A \ sets M1" "B \ sets M2" + moreover have "(fst \ f) -` A \ space M \ sets M" "(snd \ f) -` B \ space M \ sets M" + using f `A \ sets M1` s `B \ sets M2` by (auto simp: measurable_sets) + moreover have "f -` (A \ B) \ space M = ((fst \ f) -` A \ space M) \ ((snd \ f) -` B \ space M)" + by (auto simp: vimage_Times) + ultimately show "f -` (A \ B) \ space M \ sets M" by auto qed -qed +qed auto -lemma (in pair_sigma_algebra) measurable_pair: - assumes "sigma_algebra M" - assumes "(fst \ f) \ measurable M M1" "(snd \ f) \ measurable M M2" - shows "f \ measurable M P" - unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp - -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_measure_generator_def by auto +lemma measurable_pair: + "(fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2 \ f \ measurable M (M1 \\<^isub>M M2)" + unfolding measurable_pair_iff by simp -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_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_measure_generatorI) +lemma measurable_pair_swap': "(\(x,y). (y, x)) \ measurable (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" +proof (rule measurable_pair_measureI) + fix A B assume "A \ sets M2" "B \ sets M1" + moreover then have "(\(x, y). (y, x)) -` (A \ B) \ space (M1 \\<^isub>M M2) = B \ A" + by (auto dest: sets_into_space simp: space_pair_measure) + ultimately show "(\(x, y). (y, x)) -` (A \ B) \ space (M1 \\<^isub>M M2) \ sets (M1 \\<^isub>M M2)" + by auto +qed (auto simp add: space_pair_measure) + +lemma measurable_pair_swap: + assumes f: "f \ measurable (M1 \\<^isub>M M2) M" shows "(\(x,y). f (y, x)) \ measurable (M2 \\<^isub>M M1) M" +proof - + have "(\x. f (case x of (x, y) \ (y, x))) = (\(x, y). f (y, x))" by auto + then show ?thesis + using measurable_comp[OF measurable_pair_swap' f] by (simp add: comp_def) qed -lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap: - assumes Q: "Q \ sets P" - shows "(\(x,y). (y, x)) -` Q \ sets (M2 \\<^isub>M M1)" (is "_ \ sets ?Q") -proof - - 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 measurable_pair_swap_iff: + "f \ measurable (M2 \\<^isub>M M1) M \ (\(x,y). f (y,x)) \ measurable (M1 \\<^isub>M M2) M" + using measurable_pair_swap[of "\(x,y). f (y, x)"] + by (auto intro!: measurable_pair_swap) -lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable: - 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_measure_generator_def pair_measure_def by auto -next - fix A assume "A \ sets (M2 \\<^isub>M M1)" - interpret Q: pair_sigma_algebra M2 M1 by default - from Q.sets_pair_sigma_algebra_swap[OF `A \ sets (M2 \\<^isub>M M1)`] - show "?f -` A \ space ?P \ sets ?P" by simp -qed +lemma measurable_Pair1': "x \ space M1 \ Pair x \ measurable M2 (M1 \\<^isub>M M2)" +proof (rule measurable_pair_measureI) + fix A B assume "A \ sets M1" "B \ sets M2" + moreover then have "Pair x -` (A \ B) \ space M2 = (if x \ A then B else {})" + by (auto dest: sets_into_space simp: space_pair_measure) + ultimately show "Pair x -` (A \ B) \ space M2 \ sets M2" + by auto +qed (auto simp add: space_pair_measure) -lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]: - assumes "Q \ sets P" shows "Pair x -` Q \ sets M2" +lemma sets_Pair1: assumes A: "A \ sets (M1 \\<^isub>M M2)" shows "Pair x -` A \ sets M2" proof - - 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_measure) - have "sets E \ sets ?Q" - using M1.sets_into_space M2.sets_into_space - by (auto simp: pair_measure_generator_def space_pair_measure) - then have "sets P \ sets ?Q" - apply (subst pair_measure_def, intro Q.sets_sigma_subset) - by (simp add: pair_measure_def) - with assms show ?thesis by auto + have "Pair x -` A = (if x \ space M1 then Pair x -` A \ space M2 else {})" + using A[THEN sets_into_space] by (auto simp: space_pair_measure) + also have "\ \ sets M2" + using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm) + finally show ?thesis . qed -lemma (in pair_sigma_algebra) measurable_cut_snd: - 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 - from Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y] - show ?thesis by (simp add: vimage_compose[symmetric] comp_def) -qed +lemma measurable_Pair2': "y \ space M2 \ (\x. (x, y)) \ measurable M1 (M1 \\<^isub>M M2)" + using measurable_comp[OF measurable_Pair1' measurable_pair_swap', of y M2 M1] + by (simp add: comp_def) -lemma (in pair_sigma_algebra) measurable_pair_image_snd: - assumes m: "f \ measurable P M" and "x \ space M1" - shows "(\y. f (x, y)) \ measurable M2 M" - 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_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_measure_generator_def pair_measure_def) - finally show "(\y. f (x, y)) -` A \ space M2 \ sets M2" . +lemma sets_Pair2: assumes A: "A \ sets (M1 \\<^isub>M M2)" shows "(\x. (x, y)) -` A \ sets M1" +proof - + have "(\x. (x, y)) -` A = (if y \ space M2 then (\x. (x, y)) -` A \ space M1 else {})" + using A[THEN sets_into_space] by (auto simp: space_pair_measure) + also have "\ \ sets M1" + using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm) + finally show ?thesis . qed -lemma (in pair_sigma_algebra) measurable_pair_image_fst: - assumes m: "f \ measurable P M" and "y \ space M2" +lemma measurable_Pair2: + assumes f: "f \ measurable (M1 \\<^isub>M M2) M" and x: "x \ space M1" + shows "(\y. f (x, y)) \ measurable M2 M" + using measurable_comp[OF measurable_Pair1' f, OF x] + by (simp add: comp_def) + +lemma measurable_Pair1: + assumes f: "f \ measurable (M1 \\<^isub>M M2) M" and y: "y \ space M2" shows "(\x. f (x, y)) \ measurable M1 M" -proof - - interpret Q: pair_sigma_algebra M2 M1 by default - from Q.measurable_pair_image_snd[OF measurable_comp `y \ space M2`, - OF Q.pair_sigma_algebra_swap_measurable m] - show ?thesis by simp -qed + using measurable_comp[OF measurable_Pair2' f, OF y] + by (simp add: comp_def) -lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E" +lemma Int_stable_pair_measure_generator: "Int_stable {a \ b | a b. a \ sets A \ b \ sets B}" 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_measure_generator_def by auto - then show "A \ B \ sets E" - by (auto simp add: times_Int_times pair_measure_generator_def) -qed + by safe (auto simp add: times_Int_times) lemma finite_measure_cut_measurable: - 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" + shows "(\x. emeasure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") proof - 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} \" + let ?\ = "space M1 \ space M2" and ?D = "{A\sets (M1 \\<^isub>M M2). ?s A \ borel_measurable M1}" note space_pair_measure[simp] - interpret dynkin_system ?D + interpret dynkin_system ?\ ?D proof (intro dynkin_systemI) - fix A assume "A \ sets ?D" then show "A \ space ?D" - using sets_into_space by simp + fix A assume "A \ ?D" then show "A \ ?\" + using sets_into_space[of A "M1 \\<^isub>M M2"] by simp next - from top show "space ?D \ sets ?D" - by (auto simp add: if_distrib intro!: M1.measurable_If) + from top show "?\ \ ?D" + by (auto simp add: if_distrib intro!: measurable_If) next - fix A assume "A \ sets ?D" - 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.measure_compl simp: vimage_Diff) - with `A \ sets ?D` top show "space ?D - A \ sets ?D" - by (auto intro!: Diff M1.measurable_If M1.borel_measurable_ereal_diff) + fix A assume "A \ ?D" + with sets_into_space have "\x. emeasure M2 (Pair x -` (?\ - A)) = + (if x \ space M1 then emeasure M2 (space M2) - ?s A x else 0)" + by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) + with `A \ ?D` top show "?\ - A \ ?D" + by (auto intro!: measurable_If) next - fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ sets ?D" - moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\i. ?s (F i) x)" - by (intro M2.measure_countably_additive[symmetric]) - (auto simp: disjoint_family_on_def) - ultimately show "(\i. F i) \ sets ?D" - by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf) + fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ ?D" + moreover then have "\x. emeasure M2 (\i. Pair x -` F i) = (\i. ?s (F i) x)" + by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1) + ultimately show "(\i. F i) \ ?D" + by (auto simp: vimage_UN intro!: borel_measurable_psuminf) qed - have "sets P = sets ?D" apply (subst pair_measure_def) - proof (intro dynkin_lemma) - 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_measure_generator_def sets_sigma if_distrib - intro: sigma_sets.Basic intro!: M1.measurable_If) - qed (auto simp: pair_measure_def) - with `Q \ sets P` have "Q \ sets ?D" by simp + let ?G = "{a \ b | a b. a \ sets M1 \ b \ sets M2}" + have "sigma_sets ?\ ?G = ?D" + proof (rule dynkin_lemma) + show "?G \ ?D" + by (auto simp: if_distrib Int_def[symmetric] intro!: measurable_If) + qed (auto simp: sets_pair_measure Int_stable_pair_measure_generator) + with `Q \ sets (M1 \\<^isub>M M2)` have "Q \ ?D" + by (simp add: sets_pair_measure[symmetric]) then show "?s Q \ borel_measurable M1" by simp qed -subsection {* Binary products of $\sigma$-finite measure spaces *} +subsection {* Binary products of $\sigma$-finite emeasure spaces *} -locale pair_sigma_finite = pair_sigma_algebra M1 M2 + M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 - for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme" +locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 + for M1 :: "'a measure" and M2 :: "'b measure" -lemma (in pair_sigma_finite) measure_cut_measurable_fst: - assumes "Q \ sets P" shows "(\x. measure M2 (Pair x -` Q)) \ borel_measurable M1" (is "?s Q \ _") +lemma sets_pair_measure_cong[cong]: + "sets M1 = sets M1' \ sets M2 = sets M2' \ sets (M1 \\<^isub>M M2) = sets (M1' \\<^isub>M M2')" + unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq) + +lemma (in pair_sigma_finite) measurable_emeasure_Pair1: + assumes Q: "Q \ sets (M1 \\<^isub>M M2)" shows "(\x. emeasure 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" by default - from M2.disjoint_sigma_finite guess F .. note F = this + from M2.sigma_finite_disjoint guess F . note F = this then have F_sets: "\i. F i \ sets M2" by auto + have M1: "sigma_finite_measure M1" .. 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 - let ?R2 = "M2.restricted_space (F i)" - have "(\x. measure ?R2 (Pair x -` (space M1 \ space ?R2 \ Q))) \ borel_measurable M1" + using F sets_into_space by auto + let ?R = "density M2 (indicator (F i))" + have "(\x. emeasure ?R (Pair x -` (space M1 \ space ?R \ Q))) \ borel_measurable M1" proof (intro finite_measure_cut_measurable[OF M1]) - show "finite_measure ?R2" - using F by (intro M2.restricted_to_finite_measure) auto - 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)" . + show "finite_measure ?R" + using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) + show "(space M1 \ space ?R) \ Q \ sets (M1 \\<^isub>M ?R)" + using Q by (simp add: Int) 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_measure) - ultimately have "(\x. measure M2 (?C x i)) \ borel_measurable M1" + moreover have "\x. emeasure ?R (Pair x -` (space M1 \ space ?R \ Q)) + = emeasure M2 (F i \ Pair x -` (space M1 \ space ?R \ Q))" + using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) + moreover have "\x. F i \ Pair x -` (space M1 \ space ?R \ Q) = ?C x i" + using sets_into_space[OF Q] by (auto simp: space_pair_measure) + ultimately have "(\x. emeasure M2 (?C x i)) \ borel_measurable M1" by simp } moreover { fix x - have "(\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" - proof (intro M2.measure_countably_additive) + have "(\i. emeasure M2 (?C x i)) = emeasure M2 (\i. ?C x i)" + proof (intro suminf_emeasure) show "range (?C x) \ sets M2" - using F `Q \ sets P` by (auto intro!: M2.Int) + using F `Q \ sets (M1 \\<^isub>M M2)` by (auto intro!: sets_Pair1) have "disjoint_family F" using F by auto show "disjoint_family (?C x)" by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto qed also have "(\i. ?C x i) = Pair x -` Q" - using F sets_into_space `Q \ sets P` + using F sets_into_space[OF `Q \ sets (M1 \\<^isub>M M2)`] by (auto simp: space_pair_measure) - finally have "measure M2 (Pair x -` Q) = (\i. measure M2 (?C x i))" + finally have "emeasure M2 (Pair x -` Q) = (\i. emeasure M2 (?C x i))" by simp } - ultimately show ?thesis using `Q \ sets P` F_sets - by (auto intro!: M1.borel_measurable_psuminf M2.Int) + ultimately show ?thesis using `Q \ sets (M1 \\<^isub>M M2)` F_sets + by auto qed -lemma (in pair_sigma_finite) measure_cut_measurable_snd: - assumes "Q \ sets P" shows "(\y. M1.\ ((\x. (x, y)) -` Q)) \ borel_measurable M2" +lemma (in pair_sigma_finite) measurable_emeasure_Pair2: + assumes Q: "Q \ sets (M1 \\<^isub>M M2)" shows "(\y. emeasure M1 ((\x. (x, y)) -` Q)) \ borel_measurable M2" proof - 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 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 (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) - show ?thesis - using Q.pair_sigma_algebra_swap_measurable assms - unfolding * by (rule measurable_comp) + have "(\(x, y). (y, x)) -` Q \ space (M2 \\<^isub>M M1) \ sets (M2 \\<^isub>M M1)" + using Q measurable_pair_swap' by (auto intro: measurable_sets) + note Q.measurable_emeasure_Pair1[OF this] + moreover have "\y. Pair y -` ((\(x, y). (y, x)) -` Q \ space (M2 \\<^isub>M M1)) = (\x. (x, y)) -` Q" + using Q[THEN sets_into_space] by (auto simp: space_pair_measure) + ultimately show ?thesis by simp qed -lemma (in pair_sigma_finite) pair_measure_alt: - assumes "A \ sets P" - 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 :: ereal)" - unfolding indicator_def by auto - 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]) - by simp +lemma (in pair_sigma_finite) emeasure_pair_measure: + assumes "X \ sets (M1 \\<^isub>M M2)" + shows "emeasure (M1 \\<^isub>M M2) X = (\\<^isup>+ x. \\<^isup>+ y. indicator X (x, y) \M2 \M1)" (is "_ = ?\ X") +proof (rule emeasure_measure_of[OF pair_measure_def]) + show "positive (sets (M1 \\<^isub>M M2)) ?\" + by (auto simp: positive_def positive_integral_positive) + have eq[simp]: "\A x y. indicator A (x, y) = indicator (Pair x -` A) y" + by (auto simp: indicator_def) + show "countably_additive (sets (M1 \\<^isub>M M2)) ?\" + proof (rule countably_additiveI) + fix F :: "nat \ ('a \ 'b) set" assume F: "range F \ sets (M1 \\<^isub>M M2)" "disjoint_family F" + from F have *: "\i. F i \ sets (M1 \\<^isub>M M2)" "(\i. F i) \ sets (M1 \\<^isub>M M2)" by auto + moreover from F have "\i. (\x. emeasure M2 (Pair x -` F i)) \ borel_measurable M1" + by (intro measurable_emeasure_Pair1) auto + moreover have "\x. disjoint_family (\i. Pair x -` F i)" + by (intro disjoint_family_on_bisimulation[OF F(2)]) auto + moreover have "\x. range (\i. Pair x -` F i) \ sets M2" + using F by (auto simp: sets_Pair1) + ultimately show "(\n. ?\ (F n)) = ?\ (\i. F i)" + by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1 + intro!: positive_integral_cong positive_integral_indicator[symmetric]) + qed + show "{a \ b |a b. a \ sets M1 \ b \ sets M2} \ Pow (space M1 \ space M2)" + using space_closed[of M1] space_closed[of M2] by auto +qed fact + +lemma (in pair_sigma_finite) emeasure_pair_measure_alt: + assumes X: "X \ sets (M1 \\<^isub>M M2)" + shows "emeasure (M1 \\<^isub>M M2) X = (\\<^isup>+x. emeasure M2 (Pair x -` X) \M1)" +proof - + have [simp]: "\x y. indicator X (x, y) = indicator (Pair x -` X) y" + by (auto simp: indicator_def) + show ?thesis + using X by (auto intro!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1) qed -lemma (in pair_sigma_finite) pair_measure_times: - assumes A: "A \ sets M1" and "B \ sets M2" - shows "measure (M1 \\<^isub>M M2) (A \ B) = M1.\ A * measure M2 B" +lemma (in pair_sigma_finite) emeasure_pair_measure_Times: + assumes A: "A \ sets M1" and B: "B \ sets M2" + shows "emeasure (M1 \\<^isub>M M2) (A \ B) = emeasure M1 A * emeasure M2 B" proof - - 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) + have "emeasure (M1 \\<^isub>M M2) (A \ B) = (\\<^isup>+x. emeasure M2 B * indicator A x \M1)" + using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt) + also have "\ = emeasure M2 B * emeasure M1 A" + using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator) + finally show ?thesis + by (simp add: ac_simps) qed -lemma (in measure_space) measure_not_negative[simp,intro]: - assumes A: "A \ sets M" shows "\ A \ - \" - using positive_measure[OF A] by auto - lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: - "\F::nat \ ('a \ 'b) set. range F \ sets E \ incseq F \ (\i. F i) = space E \ - (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" + defines "E \ {A \ B | A B. A \ sets M1 \ B \ sets M2}" + shows "\F::nat \ ('a \ 'b) set. range F \ E \ incseq F \ (\i. F i) = space M1 \ space M2 \ + (\i. emeasure (M1 \\<^isub>M M2) (F i) \ \)" proof - - obtain F1 :: "nat \ 'a set" and F2 :: "nat \ 'b set" where - F1: "range F1 \ sets M1" "incseq F1" "(\i. F1 i) = space M1" "\i. M1.\ (F1 i) \ \" and - F2: "range F2 \ sets M2" "incseq F2" "(\i. F2 i) = 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)" by auto + from M1.sigma_finite_incseq guess F1 . note F1 = this + from M2.sigma_finite_incseq guess F2 . note F2 = this + from F1 F2 have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" by auto let ?F = "\i. F1 i \ F2 i" - show ?thesis unfolding space_pair_measure + show ?thesis proof (intro exI[of _ ?F] conjI allI) - show "range ?F \ sets E" using F1 F2 - by (fastforce intro!: pair_measure_generatorI) + show "range ?F \ E" using F1 F2 by (auto simp: E_def) (metis range_subsetD) next have "space M1 \ space M2 \ (\i. ?F i)" proof (intro subsetI) @@ -448,353 +352,315 @@ 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 E" - using space by (auto simp: space pair_measure_generator_def) + then show "(\i. ?F i) = space M1 \ space M2" + using space by (auto simp: space) next fix i show "incseq (\i. F1 i \ F2 i)" using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto next fix i from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto - with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)] - show "measure P (F1 i \ F2 i) \ \" - by (simp add: pair_measure_times) + with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"] + show "emeasure (M1 \\<^isub>M M2) (F1 i \ F2 i) \ \" + by (auto simp add: emeasure_pair_measure_Times) + qed +qed + +sublocale pair_sigma_finite \ sigma_finite_measure "M1 \\<^isub>M M2" +proof + 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 (M1 \\<^isub>M M2) \ (\i. F i) = space (M1 \\<^isub>M M2) \ (\i. emeasure (M1 \\<^isub>M M2) (F i) \ \)" + proof (rule exI[of _ F], intro conjI) + show "range F \ sets (M1 \\<^isub>M M2)" using F by (auto simp: pair_measure_def) + show "(\i. F i) = space (M1 \\<^isub>M M2)" + using F by (auto simp: space_pair_measure) + show "\i. emeasure (M1 \\<^isub>M M2) (F i) \ \" using F by auto qed qed -sublocale pair_sigma_finite \ sigma_finite_measure P -proof - show "positive P (measure P)" - unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def - by (auto intro: M1.positive_integral_positive M2.positive_integral_positive) - - 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. 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 - ultimately show "(\n. measure P (F n)) = measure P (\i. F i)" - by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric] - M2.measure_countably_additive - cong: M1.positive_integral_cong) - qed - - 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 simp: pair_measure_def) - show "(\i. F i) = space P" - using F by (auto simp: pair_measure_def pair_measure_generator_def) - show "\i. measure P (F i) \ \" using F by auto - qed +lemma sigma_finite_pair_measure: + assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B" + shows "sigma_finite_measure (A \\<^isub>M B)" +proof - + interpret A: sigma_finite_measure A by fact + interpret B: sigma_finite_measure B by fact + interpret AB: pair_sigma_finite A B .. + show ?thesis .. qed -lemma (in pair_sigma_algebra) sets_swap: - assumes "A \ sets P" +lemma sets_pair_swap: + assumes "A \ sets (M1 \\<^isub>M M2)" 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_measure) - show ?thesis - unfolding * using assms by (rule sets_pair_sigma_algebra_swap) -qed + using measurable_pair_swap' assms by (rule measurable_sets) -lemma (in pair_sigma_finite) pair_measure_alt2: - assumes A: "A \ sets P" - shows "\ A = (\\<^isup>+y. M1.\ ((\x. (x, y)) -` A) \M2)" - (is "_ = ?\ A") +lemma (in pair_sigma_finite) distr_pair_swap: + "M1 \\<^isub>M M2 = distr (M2 \\<^isub>M M1) (M1 \\<^isub>M M2) (\(x, y). (y, x))" (is "?P = ?D") proof - interpret Q: pair_sigma_finite M2 M1 by default 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 - - have "\ A = Q.\ ((\(y, x). (x, y)) -` A \ space Q.P)" - proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator]) - show "measure_space P" "measure_space Q.P" by default - show "(\(y, x). (x, y)) \ measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable) - show "sets (sigma E) = sets P" "space E = space P" "A \ sets (sigma E)" - using assms unfolding pair_measure_def by auto - show "range F \ sets E" "incseq F" "(\i. F i) = space E" "\i. \ (F i) \ \" - using F `A \ sets P` by (auto simp: pair_measure_def) - fix X assume "X \ sets E" - then obtain A B where X[simp]: "X = A \ B" and AB: "A \ sets M1" "B \ sets M2" - unfolding pair_measure_def pair_measure_generator_def by auto - then have "(\(y, x). (x, y)) -` X \ space Q.P = B \ A" - using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure) - then show "\ X = Q.\ ((\(y, x). (x, y)) -` X \ space Q.P)" - using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps) + let ?E = "{a \ b |a b. a \ sets M1 \ b \ sets M2}" + show ?thesis + proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) + show "?E \ Pow (space ?P)" + using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure) + show "sets ?P = sigma_sets (space ?P) ?E" + by (simp add: sets_pair_measure space_pair_measure) + then show "sets ?D = sigma_sets (space ?P) ?E" + by simp + next + show "range F \ ?E" "incseq F" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" + using F by (auto simp: space_pair_measure) + next + fix X assume "X \ ?E" + then obtain A B where X[simp]: "X = A \ B" and A: "A \ sets M1" and B: "B \ sets M2" by auto + have "(\(y, x). (x, y)) -` X \ space (M2 \\<^isub>M M1) = B \ A" + using sets_into_space[OF A] sets_into_space[OF B] by (auto simp: space_pair_measure) + with A B show "emeasure (M1 \\<^isub>M M2) X = emeasure ?D X" + by (simp add: emeasure_pair_measure_Times Q.emeasure_pair_measure_Times emeasure_distr + measurable_pair_swap' ac_simps) qed - then show ?thesis - using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]] - by (auto simp add: Q.pair_measure_alt space_pair_measure - intro!: M2.positive_integral_cong arg_cong[where f="M1.\"]) qed -lemma pair_sigma_algebra_sigma: - assumes 1: "incseq S1" "(\i. S1 i) = space E1" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" - assumes 2: "decseq S2" "(\i. S2 i) = 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") +lemma (in pair_sigma_finite) emeasure_pair_measure_alt2: + assumes A: "A \ sets (M1 \\<^isub>M M2)" + shows "emeasure (M1 \\<^isub>M M2) A = (\\<^isup>+y. emeasure M1 ((\x. (x, y)) -` A) \M2)" + (is "_ = ?\ A") 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 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) - (force 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 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) - (force 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 + interpret Q: pair_sigma_finite M2 M1 by default + have [simp]: "\y. (Pair y -` ((\(x, y). (y, x)) -` A \ space (M2 \\<^isub>M M1))) = (\x. (x, y)) -` A" + using sets_into_space[OF A] by (auto simp: space_pair_measure) + show ?thesis using A + by (subst distr_pair_swap) + (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap'] + Q.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A]) qed section "Fubinis theorem" lemma (in pair_sigma_finite) simple_function_cut: - assumes f: "simple_function P f" "\x. 0 \ f x" + assumes f: "simple_function (M1 \\<^isub>M M2) f" "\x. 0 \ f x" 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" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P (M1 \\<^isub>M M2) f" proof - - have f_borel: "f \ borel_measurable P" + have f_borel: "f \ borel_measurable (M1 \\<^isub>M M2)" using f(1) by (rule borel_measurable_simple_function) - let ?F = "\z. f -` {z} \ space P" + let ?F = "\z. f -` {z} \ space (M1 \\<^isub>M M2)" let ?F' = "\x z. Pair x -` ?F z" { fix x assume "x \ space M1" 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` + have "\y. y \ space M2 \ (x, y) \ space (M1 \\<^isub>M M2)" using `x \ space M1` 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) + by (rule sets_Pair1[OF measurable_sets]) auto ultimately have "simple_function M2 (\ y. f (x, y))" - apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _]) + apply (rule_tac simple_function_cong[THEN iffD2, OF _]) apply (rule simple_function_indicator_representation[OF f(1)]) - using `x \ space M1` by (auto simp del: space_sigma) } + using `x \ space M1` by auto } note M2_sf = this { fix x assume x: "x \ space M1" - 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] f(2)] + then have "(\\<^isup>+y. f (x, y) \M2) = (\z\f ` space (M1 \\<^isub>M M2). z * emeasure M2 (?F' x z))" + unfolding positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)] unfolding simple_integral_def proof (safe intro!: setsum_mono_zero_cong_left) - from f(1) show "finite (f ` space P)" by (rule simple_functionD) + from f(1) show "finite (f ` space (M1 \\<^isub>M M2))" by (rule simple_functionD) next - fix y assume "y \ space M2" then show "f (x, y) \ f ` space P" + fix y assume "y \ space M2" then show "f (x, y) \ f ` space (M1 \\<^isub>M M2)" using `x \ space M1` by (auto simp: space_pair_measure) next - fix x' y assume "(x', y) \ space P" + fix x' y assume "(x', y) \ space (M1 \\<^isub>M M2)" "f (x', y) \ (\y. f (x, y)) ` space M2" then have *: "?F' x (f (x', y)) = {}" by (force simp: space_pair_measure) - show "f (x', y) * M2.\ (?F' x (f (x', y))) = 0" + show "f (x', y) * emeasure M2 (?F' x (f (x', y))) = 0" unfolding * by simp qed (simp add: vimage_compose[symmetric] comp_def 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. M2.\ (?F' x z)) \ borel_measurable M1" - by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int) - moreover have *: "\i x. 0 \ M2.\ (Pair x -` (f -` {i} \ space P))" - using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst) - moreover { fix i assume "i \ f`space P" - with * have "\x. 0 \ i * M2.\ (Pair x -` (f -` {i} \ space P))" + moreover have "\z. ?F z \ sets (M1 \\<^isub>M M2)" + by (auto intro!: f_borel borel_measurable_vimage) + moreover then have "\z. (\x. emeasure M2 (?F' x z)) \ borel_measurable M1" + by (auto intro!: measurable_emeasure_Pair1 simp del: vimage_Int) + moreover have *: "\i x. 0 \ emeasure M2 (Pair x -` (f -` {i} \ space (M1 \\<^isub>M M2)))" + using f(1)[THEN simple_functionD(2)] f(2) by (intro emeasure_nonneg) + moreover { fix i assume "i \ f`space (M1 \\<^isub>M M2)" + with * have "\x. 0 \ i * emeasure M2 (Pair x -` (f -` {i} \ space (M1 \\<^isub>M M2)))" using f(2) by auto } ultimately 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" using f(2) - by (auto simp del: vimage_Int cong: measurable_cong - intro!: M1.borel_measurable_ereal_setsum setsum_cong - simp add: M1.positive_integral_setsum simple_integral_def - M1.positive_integral_cmult - M1.positive_integral_cong[OF eq] + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P (M1 \\<^isub>M M2) f" using f(2) + by (auto simp del: vimage_Int cong: measurable_cong intro!: setsum_cong + simp add: positive_integral_setsum simple_integral_def + positive_integral_cmult + positive_integral_cong[OF eq] positive_integral_eq_simple_integral[OF f] - pair_measure_alt[symmetric]) + emeasure_pair_measure_alt[symmetric]) qed lemma (in pair_sigma_finite) positive_integral_fst_measurable: - assumes f: "f \ borel_measurable P" + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" shows "(\x. \\<^isup>+ y. f (x, y) \M2) \ borel_measurable M1" (is "?C f \ borel_measurable M1") - and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P (M1 \\<^isub>M M2) f" proof - from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this - then have F_borel: "\i. F i \ borel_measurable P" + then have F_borel: "\i. F i \ borel_measurable (M1 \\<^isub>M M2)" by (auto intro: borel_measurable_simple_function) note sf = simple_function_cut[OF F(1,5)] then have "(\x. SUP i. ?C (F i) x) \ borel_measurable M1" using F(1) by auto moreover { fix x assume "x \ space M1" - from F measurable_pair_image_snd[OF F_borel`x \ space M1`] + from F measurable_Pair2[OF F_borel `x \ space M1`] have "(\\<^isup>+y. (SUP i. F i (x, y)) \M2) = (SUP i. ?C (F i) x)" - by (intro M2.positive_integral_monotone_convergence_SUP) + by (intro positive_integral_monotone_convergence_SUP) (auto simp: incseq_Suc_iff le_fun_def) then have "(SUP i. ?C (F i) x) = ?C f x" unfolding F(4) positive_integral_max_0 by simp } note SUPR_C = this ultimately show "?C f \ borel_measurable M1" by (simp cong: measurable_cong) - have "(\\<^isup>+x. (SUP i. F i x) \P) = (SUP i. integral\<^isup>P P (F i))" + have "(\\<^isup>+x. (SUP i. F i x) \(M1 \\<^isub>M M2)) = (SUP i. integral\<^isup>P (M1 \\<^isub>M M2) (F i))" using F_borel F by (intro positive_integral_monotone_convergence_SUP) auto - also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" + also have "(SUP i. integral\<^isup>P (M1 \\<^isub>M M2) (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" unfolding sf(2) by simp also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" using F sf(1) - by (intro M1.positive_integral_monotone_convergence_SUP[symmetric]) - (auto intro!: M2.positive_integral_mono M2.positive_integral_positive - simp: incseq_Suc_iff le_fun_def) + by (intro positive_integral_monotone_convergence_SUP[symmetric]) + (auto intro!: positive_integral_mono positive_integral_positive + simp: incseq_Suc_iff le_fun_def) also have "\ = \\<^isup>+ x. (\\<^isup>+ y. (SUP i. F i (x, y)) \M2) \M1" using F_borel F(2,5) - by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric] - simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd) - finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" + by (auto intro!: positive_integral_cong positive_integral_monotone_convergence_SUP[symmetric] measurable_Pair2 + simp: incseq_Suc_iff le_fun_def) + finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P (M1 \\<^isub>M M2) f" using F by (simp add: positive_integral_max_0) qed -lemma (in pair_sigma_finite) measure_preserving_swap: - "(\(x,y). (y, x)) \ measure_preserving (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" -proof - interpret Q: pair_sigma_finite M2 M1 by default - show *: "(\(x,y). (y, x)) \ measurable (M1 \\<^isub>M M2) (M2 \\<^isub>M M1)" - using pair_sigma_algebra_swap_measurable . - fix X assume "X \ sets (M2 \\<^isub>M M1)" - from measurable_sets[OF * this] this Q.sets_into_space[OF this] - show "measure (M1 \\<^isub>M M2) ((\(x, y). (y, x)) -` X \ space P) = measure (M2 \\<^isub>M M1) X" - by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\"] - simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure) -qed - -lemma (in pair_sigma_finite) positive_integral_product_swap: - assumes f: "f \ borel_measurable P" - shows "(\\<^isup>+x. f (case x of (x,y)\(y,x)) \(M2 \\<^isub>M M1)) = integral\<^isup>P P f" +lemma (in pair_sigma_finite) positive_integral_snd_measurable: + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" + shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = integral\<^isup>P (M1 \\<^isub>M M2) f" proof - interpret Q: pair_sigma_finite M2 M1 by default - have "sigma_algebra P" by default - with f show ?thesis - by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto -qed - -lemma (in pair_sigma_finite) positive_integral_snd_measurable: - assumes f: "f \ borel_measurable P" - shows "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = integral\<^isup>P P f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - note pair_sigma_algebra_measurable[OF f] + note measurable_pair_swap[OF f] from Q.positive_integral_fst_measurable[OF this] - have "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ (x, y). f (y, x) \Q.P)" + have "(\\<^isup>+ y. (\\<^isup>+ x. f (x, y) \M1) \M2) = (\\<^isup>+ (x, y). f (y, x) \(M2 \\<^isub>M M1))" by simp - 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) + also have "(\\<^isup>+ (x, y). f (y, x) \(M2 \\<^isub>M M1)) = integral\<^isup>P (M1 \\<^isub>M M2) f" + by (subst distr_pair_swap) + (auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong) finally show ?thesis . qed lemma (in pair_sigma_finite) Fubini: - assumes f: "f \ borel_measurable P" + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" 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] .. lemma (in pair_sigma_finite) AE_pair: - assumes "AE x in P. Q x" + assumes "AE x in (M1 \\<^isub>M M2). Q x" shows "AE x in M1. (AE y in M2. Q (x, y))" proof - - obtain N where N: "N \ sets P" "\ N = 0" "{x\space P. \ Q x} \ N" - using assms unfolding almost_everywhere_def by auto + obtain N where N: "N \ sets (M1 \\<^isub>M M2)" "emeasure (M1 \\<^isub>M M2) N = 0" "{x\space (M1 \\<^isub>M M2). \ Q x} \ N" + using assms unfolding eventually_ae_filter by auto show ?thesis - proof (rule M1.AE_I) - from N measure_cut_measurable_fst[OF `N \ sets P`] - show "M1.\ {x\space M1. M2.\ (Pair x -` N) \ 0} = 0" - by (auto simp: pair_measure_alt M1.positive_integral_0_iff) - show "{x \ space M1. M2.\ (Pair x -` N) \ 0} \ sets M1" - by (intro M1.borel_measurable_ereal_neq_const measure_cut_measurable_fst N) - { 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 "M2.\ (Pair x -` N) = 0" by fact - show "Pair x -` N \ sets M2" by (intro measurable_cut_fst N) + proof (rule AE_I) + from N measurable_emeasure_Pair1[OF `N \ sets (M1 \\<^isub>M M2)`] + show "emeasure M1 {x\space M1. emeasure M2 (Pair x -` N) \ 0} = 0" + by (auto simp: emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg) + show "{x \ space M1. emeasure M2 (Pair x -` N) \ 0} \ sets M1" + by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N) + { fix x assume "x \ space M1" "emeasure M2 (Pair x -` N) = 0" + have "AE y in M2. Q (x, y)" + proof (rule AE_I) + show "emeasure M2 (Pair x -` N) = 0" by fact + show "Pair x -` N \ sets M2" using N(1) by (rule sets_Pair1) show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" - using N `x \ space M1` unfolding space_sigma space_pair_measure by auto + using N `x \ space M1` unfolding space_pair_measure by auto qed } - then show "{x \ space M1. \ M2.almost_everywhere (\y. Q (x, y))} \ {x \ space M1. M2.\ (Pair x -` N) \ 0}" + then show "{x \ space M1. \ (AE y in M2. Q (x, y))} \ {x \ space M1. emeasure M2 (Pair x -` N) \ 0}" by auto qed qed -lemma (in pair_sigma_algebra) measurable_product_swap: - "f \ measurable (M2 \\<^isub>M M1) M \ (\(x,y). f (y,x)) \ measurable P M" +lemma (in pair_sigma_finite) AE_pair_measure: + assumes "{x\space (M1 \\<^isub>M M2). P x} \ sets (M1 \\<^isub>M M2)" + assumes ae: "AE x in M1. AE y in M2. P (x, y)" + shows "AE x in M1 \\<^isub>M M2. P x" +proof (subst AE_iff_measurable[OF _ refl]) + show "{x\space (M1 \\<^isub>M M2). \ P x} \ sets (M1 \\<^isub>M M2)" + by (rule sets_Collect) fact + then have "emeasure (M1 \\<^isub>M M2) {x \ space (M1 \\<^isub>M M2). \ P x} = + (\\<^isup>+ x. \\<^isup>+ y. indicator {x \ space (M1 \\<^isub>M M2). \ P x} (x, y) \M2 \M1)" + by (simp add: emeasure_pair_measure) + also have "\ = (\\<^isup>+ x. \\<^isup>+ y. 0 \M2 \M1)" + using ae + apply (safe intro!: positive_integral_cong_AE) + apply (intro AE_I2) + apply (safe intro!: positive_integral_cong_AE) + apply auto + done + finally show "emeasure (M1 \\<^isub>M M2) {x \ space (M1 \\<^isub>M M2). \ P x} = 0" by simp +qed + +lemma (in pair_sigma_finite) AE_pair_iff: + "{x\space (M1 \\<^isub>M M2). P (fst x) (snd x)} \ sets (M1 \\<^isub>M M2) \ + (AE x in M1. AE y in M2. P x y) \ (AE x in (M1 \\<^isub>M M2). P (fst x) (snd x))" + using AE_pair[of "\x. P (fst x) (snd x)"] AE_pair_measure[of "\x. P (fst x) (snd x)"] by auto + +lemma AE_distr_iff: + assumes f: "f \ measurable M N" and P: "{x \ space N. P x} \ sets N" + shows "(AE x in distr M N f. P x) \ (AE x in M. P (f x))" +proof (subst (1 2) AE_iff_measurable[OF _ refl]) + from P show "{x \ space (distr M N f). \ P x} \ sets (distr M N f)" + by (auto intro!: sets_Collect_neg) + moreover + have "f -` {x \ space N. P x} \ space M = {x \ space M. P (f x)}" + using f by (auto dest: measurable_space) + then show "{x \ space M. \ P (f x)} \ sets M" + using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg) + moreover have "f -` {x\space N. \ P x} \ space M = {x \ space M. \ P (f x)}" + using f by (auto dest: measurable_space) + ultimately show "(emeasure (distr M N f) {x \ space (distr M N f). \ P x} = 0) = + (emeasure M {x \ space M. \ P (f x)} = 0)" + using f by (simp add: emeasure_distr) +qed + +lemma (in pair_sigma_finite) AE_commute: + assumes P: "{x\space (M1 \\<^isub>M M2). P (fst x) (snd x)} \ sets (M1 \\<^isub>M M2)" + shows "(AE x in M1. AE y in M2. P x y) \ (AE y in M2. AE x in M1. P x y)" proof - - interpret Q: pair_sigma_algebra M2 M1 by default - show ?thesis - using pair_sigma_algebra_measurable[of "\(x,y). f (y, x)"] - by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI) + interpret Q: pair_sigma_finite M2 M1 .. + have [simp]: "\x. (fst (case x of (x, y) \ (y, x))) = snd x" "\x. (snd (case x of (x, y) \ (y, x))) = fst x" + by auto + have "{x \ space (M2 \\<^isub>M M1). P (snd x) (fst x)} = + (\(x, y). (y, x)) -` {x \ space (M1 \\<^isub>M M2). P (fst x) (snd x)} \ space (M2 \\<^isub>M M1)" + by (auto simp: space_pair_measure) + also have "\ \ sets (M2 \\<^isub>M M1)" + by (intro sets_pair_swap P) + finally show ?thesis + apply (subst AE_pair_iff[OF P]) + apply (subst distr_pair_swap) + apply (subst AE_distr_iff[OF measurable_pair_swap' P]) + apply (subst Q.AE_pair_iff) + apply simp_all + done qed lemma (in pair_sigma_finite) integrable_product_swap: - assumes "integrable P f" + assumes "integrable (M1 \\<^isub>M M2) f" shows "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x))" proof - 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 integrable_def - apply (subst (1 2) positive_integral_product_swap) - using `integrable P f` unfolding integrable_def - by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) + by (rule integrable_distr[OF measurable_pair_swap']) + (simp add: distr_pair_swap[symmetric] assms) qed lemma (in pair_sigma_finite) integrable_product_swap_iff: - "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x)) \ integrable P f" + "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^isub>M M2) f" proof - 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] @@ -802,27 +668,25 @@ qed lemma (in pair_sigma_finite) integral_product_swap: - assumes "integrable P f" - shows "(\(x,y). f (y,x) \(M2 \\<^isub>M M1)) = integral\<^isup>L P f" + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" + shows "(\(x,y). f (y,x) \(M2 \\<^isub>M M1)) = integral\<^isup>L (M1 \\<^isub>M M2) f" proof - - 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 lebesgue_integral_def * - apply (subst (1 2) positive_integral_product_swap) - using `integrable P f` unfolding integrable_def - by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric]) + show ?thesis unfolding * + by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) qed lemma (in pair_sigma_finite) integrable_fst_measurable: - 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") + assumes f: "integrable (M1 \\<^isub>M M2) f" + shows "AE x in M1. integrable M2 (\ y. f (x, y))" (is "?AE") + and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^isup>L (M1 \\<^isub>M M2) f" (is "?INT") proof - + have f_borel: "f \ borel_measurable (M1 \\<^isub>M M2)" + using f by auto let ?pf = "\x. ereal (f x)" and ?nf = "\x. ereal (- f x)" have - borel: "?nf \ borel_measurable P""?pf \ borel_measurable P" and - int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" + borel: "?nf \ borel_measurable (M1 \\<^isub>M M2)""?pf \ borel_measurable (M1 \\<^isub>M M2)" and + int: "integral\<^isup>P (M1 \\<^isub>M M2) ?nf \ \" "integral\<^isup>P (M1 \\<^isub>M M2) ?pf \ \" using assms by auto have "(\\<^isup>+x. (\\<^isup>+y. ereal (f (x, y)) \M2) \M1) \ \" "(\\<^isup>+x. (\\<^isup>+y. ereal (- f (x, y)) \M2) \M1) \ \" @@ -831,69 +695,92 @@ with borel[THEN positive_integral_fst_measurable(1)] have AE_pos: "AE x in M1. (\\<^isup>+y. ereal (f (x, y)) \M2) \ \" "AE x in M1. (\\<^isup>+y. ereal (- f (x, y)) \M2) \ \" - by (auto intro!: M1.positive_integral_PInf_AE ) + by (auto intro!: positive_integral_PInf_AE ) then have AE: "AE x in M1. \\\<^isup>+y. ereal (f (x, y)) \M2\ \ \" "AE x in M1. \\\<^isup>+y. ereal (- f (x, y)) \M2\ \ \" - by (auto simp: M2.positive_integral_positive) + by (auto simp: positive_integral_positive) from AE_pos show ?AE using assms - by (simp add: measurable_pair_image_snd integrable_def) + by (simp add: measurable_Pair2[OF f_borel] integrable_def) { fix f have "(\\<^isup>+ x. - \\<^isup>+ y. ereal (f x y) \M2 \M1) = (\\<^isup>+x. 0 \M1)" - using M2.positive_integral_positive - by (intro M1.positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder) + using positive_integral_positive + by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder) then have "(\\<^isup>+ x. - \\<^isup>+ y. ereal (f x y) \M2 \M1) = 0" by simp } note this[simp] - { fix f assume borel: "(\x. ereal (f x)) \ borel_measurable P" - and int: "integral\<^isup>P P (\x. ereal (f x)) \ \" - and AE: "M1.almost_everywhere (\x. (\\<^isup>+y. ereal (f (x, y)) \M2) \ \)" + { fix f assume borel: "(\x. ereal (f x)) \ borel_measurable (M1 \\<^isub>M M2)" + and int: "integral\<^isup>P (M1 \\<^isub>M M2) (\x. ereal (f x)) \ \" + and AE: "AE x in M1. (\\<^isup>+y. ereal (f (x, y)) \M2) \ \" have "integrable M1 (\x. real (\\<^isup>+y. ereal (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_of_ereal positive_integral_fst_measurable) + using borel by (auto intro!: positive_integral_fst_measurable) have "(\\<^isup>+x. ereal (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. ereal (f (x, y)) \M2) \M1)" - using AE M2.positive_integral_positive - by (auto intro!: M1.positive_integral_cong_AE simp: ereal_real) + using AE positive_integral_positive[of M2] + by (auto intro!: positive_integral_cong_AE simp: ereal_real) then show "(\\<^isup>+x. ereal (?f x) \M1) \ \" using positive_integral_fst_measurable[OF borel] int by simp have "(\\<^isup>+x. ereal (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" - by (intro M1.positive_integral_cong_pos) - (simp add: M2.positive_integral_positive real_of_ereal_pos) + by (intro positive_integral_cong_pos) + (simp add: positive_integral_positive real_of_ereal_pos) then show "(\\<^isup>+x. ereal (- ?f x) \M1) \ \" by simp qed } with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] show ?INT - unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2] + unfolding lebesgue_integral_def[of "M1 \\<^isub>M M2"] lebesgue_integral_def[of M2] borel[THEN positive_integral_fst_measurable(2), symmetric] - using AE[THEN M1.integral_real] + using AE[THEN integral_real] by simp qed lemma (in pair_sigma_finite) integrable_snd_measurable: - 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") + assumes f: "integrable (M1 \\<^isub>M M2) f" + shows "AE y in M2. integrable M1 (\x. f (x, y))" (is "?AE") + and "(\y. (\x. f (x, y) \M1) \M2) = integral\<^isup>L (M1 \\<^isub>M M2) f" (is "?INT") proof - interpret Q: pair_sigma_finite M2 M1 by default - have Q_int: "integrable Q.P (\(x, y). f (y, x))" + have Q_int: "integrable (M2 \\<^isub>M M1) (\(x, y). f (y, x))" using f unfolding integrable_product_swap_iff . show ?INT using Q.integrable_fst_measurable(2)[OF Q_int] - using integral_product_swap[OF f] by simp + using integral_product_swap[of f] f by auto show ?AE using Q.integrable_fst_measurable(1)[OF Q_int] by simp qed +lemma (in pair_sigma_finite) positive_integral_fst_measurable': + assumes f: "(\x. f (fst x) (snd x)) \ borel_measurable (M1 \\<^isub>M M2)" + shows "(\x. \\<^isup>+ y. f x y \M2) \ borel_measurable M1" + using positive_integral_fst_measurable(1)[OF f] by simp + +lemma (in pair_sigma_finite) integral_fst_measurable: + "(\x. f (fst x) (snd x)) \ borel_measurable (M1 \\<^isub>M M2) \ (\x. \ y. f x y \M2) \ borel_measurable M1" + by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_fst_measurable') + +lemma (in pair_sigma_finite) positive_integral_snd_measurable': + assumes f: "(\x. f (fst x) (snd x)) \ borel_measurable (M1 \\<^isub>M M2)" + shows "(\y. \\<^isup>+ x. f x y \M1) \ borel_measurable M2" +proof - + interpret Q: pair_sigma_finite M2 M1 .. + show ?thesis + using measurable_pair_swap[OF f] + by (intro Q.positive_integral_fst_measurable') (simp add: split_beta') +qed + +lemma (in pair_sigma_finite) integral_snd_measurable: + "(\x. f (fst x) (snd x)) \ borel_measurable (M1 \\<^isub>M M2) \ (\y. \ x. f x y \M1) \ borel_measurable M2" + by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_snd_measurable') + lemma (in pair_sigma_finite) Fubini_integral: - assumes f: "integrable P f" + assumes f: "integrable (M1 \\<^isub>M M2) 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] .. -section "Products on finite spaces" +section {* Products on counting spaces, densities and distributions *} lemma sigma_sets_pair_measure_generator_finite: assumes "finite A" and "finite B" - shows "sigma_sets (A \ B) { a \ b | a b. a \ Pow A \ b \ Pow B} = Pow (A \ B)" + shows "sigma_sets (A \ B) { a \ b | a b. a \ A \ b \ B} = Pow (A \ B)" (is "sigma_sets ?prod ?sets = _") proof safe have fin: "finite (A \ B)" using assms by (rule finite_cartesian_product) @@ -904,8 +791,7 @@ case empty show ?case by (rule sigma_sets.Empty) next case (insert a x) - hence "{a} \ sigma_sets ?prod ?sets" - by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic) + hence "{a} \ sigma_sets ?prod ?sets" by auto 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 @@ -916,48 +802,142 @@ show "a \ A" and "b \ B" by auto qed -locale pair_finite_sigma_algebra = pair_sigma_algebra M1 M2 + M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 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 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)" - using M1.finite_space M2.finite_space - by (subst finite_pair_sigma_algebra) simp - show "sets P = Pow (space P)" - by (subst (1 2) finite_pair_sigma_algebra) simp +lemma pair_measure_count_space: + assumes A: "finite A" and B: "finite B" + shows "count_space A \\<^isub>M count_space B = count_space (A \ B)" (is "?P = ?C") +proof (rule measure_eqI) + interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact + interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact + interpret P: pair_sigma_finite "count_space A" "count_space B" by default + show eq: "sets ?P = sets ?C" + by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B) + fix X assume X: "X \ sets ?P" + with eq have X_subset: "X \ A \ B" by simp + with A B have fin_Pair: "\x. finite (Pair x -` X)" + by (intro finite_subset[OF _ B]) auto + have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B) + show "emeasure ?P X = emeasure ?C X" + apply (subst P.emeasure_pair_measure_alt[OF X]) + apply (subst emeasure_count_space) + using X_subset apply auto [] + apply (simp add: fin_Pair emeasure_count_space X_subset fin_X) + apply (subst positive_integral_count_space) + using A apply simp + apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric]) + apply (subst card_gt_0_iff) + apply (simp add: fin_Pair) + apply (subst card_SigmaI[symmetric]) + using A apply simp + using fin_Pair apply simp + using X_subset apply (auto intro!: arg_cong[where f=card]) + done qed -locale pair_finite_space = pair_sigma_finite M1 M2 + pair_finite_sigma_algebra M1 M2 + - M1: finite_measure_space M1 + M2: finite_measure_space M2 for M1 M2 +lemma pair_measure_density: + assumes f: "f \ borel_measurable M1" "AE x in M1. 0 \ f x" + assumes g: "g \ borel_measurable M2" "AE x in M2. 0 \ g x" + assumes "sigma_finite_measure M1" "sigma_finite_measure M2" + assumes "sigma_finite_measure (density M1 f)" "sigma_finite_measure (density M2 g)" + shows "density M1 f \\<^isub>M density M2 g = density (M1 \\<^isub>M M2) (\(x,y). f x * g y)" (is "?L = ?R") +proof (rule measure_eqI) + interpret M1: sigma_finite_measure M1 by fact + interpret M2: sigma_finite_measure M2 by fact + interpret D1: sigma_finite_measure "density M1 f" by fact + interpret D2: sigma_finite_measure "density M2 g" by fact + interpret L: pair_sigma_finite "density M1 f" "density M2 g" .. + interpret R: pair_sigma_finite M1 M2 .. + + fix A assume A: "A \ sets ?L" + then have indicator_eq: "\x y. indicator A (x, y) = indicator (Pair x -` A) y" + and Pair_A: "\x. Pair x -` A \ sets M2" + by (auto simp: indicator_def sets_Pair1) + have f_fst: "(\p. f (fst p)) \ borel_measurable (M1 \\<^isub>M M2)" + using measurable_comp[OF measurable_fst f(1)] by (simp add: comp_def) + have g_snd: "(\p. g (snd p)) \ borel_measurable (M1 \\<^isub>M M2)" + using measurable_comp[OF measurable_snd g(1)] by (simp add: comp_def) + have "(\x. \\<^isup>+ y. g (snd (x, y)) * indicator A (x, y) \M2) \ borel_measurable M1" + using g_snd Pair_A A by (intro R.positive_integral_fst_measurable) auto + then have int_g: "(\x. \\<^isup>+ y. g y * indicator A (x, y) \M2) \ borel_measurable M1" + by simp -lemma (in pair_finite_space) pair_measure_Pair[simp]: - assumes "a \ space M1" "b \ space M2" - shows "\ {(a, b)} = M1.\ {a} * M2.\ {b}" + show "emeasure ?L A = emeasure ?R A" + apply (subst L.emeasure_pair_measure[OF A]) + apply (subst emeasure_density) + using f_fst g_snd apply (simp add: split_beta') + using A apply simp + apply (subst positive_integral_density[OF g]) + apply (simp add: indicator_eq Pair_A) + apply (subst positive_integral_density[OF f]) + apply (rule int_g) + apply (subst R.positive_integral_fst_measurable(2)[symmetric]) + using f g A Pair_A f_fst g_snd + apply (auto intro!: positive_integral_cong_AE R.measurable_emeasure_Pair1 + simp: positive_integral_cmult indicator_eq split_beta') + apply (intro AE_I2 impI) + apply (subst mult_assoc) + apply (subst positive_integral_cmult) + apply auto + done +qed simp + +lemma sigma_finite_measure_distr: + assumes "sigma_finite_measure (distr M N f)" and f: "f \ measurable M N" + shows "sigma_finite_measure M" proof - - 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 + interpret sigma_finite_measure "distr M N f" by fact + from sigma_finite_disjoint guess A . note A = this + show ?thesis + proof (unfold_locales, intro conjI exI allI) + show "range (\i. f -` A i \ space M) \ sets M" + using A f by (auto intro!: measurable_sets) + show "(\i. f -` A i \ space M) = space M" + using A(1) A(2)[symmetric] f by (auto simp: measurable_def Pi_def) + fix i show "emeasure M (f -` A i \ space M) \ \" + using f A(1,2) A(3)[of i] by (simp add: emeasure_distr subset_eq) + qed qed -lemma (in pair_finite_space) pair_measure_singleton[simp]: - assumes "x \ space M1 \ space M2" - shows "\ {x} = M1.\ {fst x} * M2.\ {snd x}" - using pair_measure_Pair assms by (cases x) auto +lemma measurable_cong': + assumes sets: "sets M = sets M'" "sets N = sets N'" + shows "measurable M N = measurable M' N'" + using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) -sublocale pair_finite_space \ finite_measure_space P -proof unfold_locales - show "measure P (space P) \ \" - by (subst (2) finite_pair_sigma_algebra) - (simp add: pair_measure_times) +lemma pair_measure_distr: + assumes f: "f \ measurable M S" and g: "g \ measurable N T" + assumes "sigma_finite_measure (distr M S f)" "sigma_finite_measure (distr N T g)" + shows "distr M S f \\<^isub>M distr N T g = distr (M \\<^isub>M N) (S \\<^isub>M T) (\(x, y). (f x, g y))" (is "?P = ?D") +proof (rule measure_eqI) + show "sets ?P = sets ?D" + by simp + interpret S: sigma_finite_measure "distr M S f" by fact + interpret T: sigma_finite_measure "distr N T g" by fact + interpret ST: pair_sigma_finite "distr M S f" "distr N T g" .. + interpret M: sigma_finite_measure M by (rule sigma_finite_measure_distr) fact+ + interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+ + interpret MN: pair_sigma_finite M N .. + interpret SN: pair_sigma_finite "distr M S f" N .. + have [simp]: + "\f g. fst \ (\(x, y). (f x, g y)) = f \ fst" "\f g. snd \ (\(x, y). (f x, g y)) = g \ snd" + by auto + then have fg: "(\(x, y). (f x, g y)) \ measurable (M \\<^isub>M N) (S \\<^isub>M T)" + using measurable_comp[OF measurable_fst f] measurable_comp[OF measurable_snd g] + by (auto simp: measurable_pair_iff) + fix A assume A: "A \ sets ?P" + then have "emeasure ?P A = (\\<^isup>+x. emeasure (distr N T g) (Pair x -` A) \distr M S f)" + by (rule ST.emeasure_pair_measure_alt) + also have "\ = (\\<^isup>+x. emeasure N (g -` (Pair x -` A) \ space N) \distr M S f)" + using g A by (simp add: sets_Pair1 emeasure_distr) + also have "\ = (\\<^isup>+x. emeasure N (g -` (Pair (f x) -` A) \ space N) \M)" + using f g A ST.measurable_emeasure_Pair1[OF A] + by (intro positive_integral_distr) (auto simp add: sets_Pair1 emeasure_distr) + also have "\ = (\\<^isup>+x. emeasure N (Pair x -` ((\(x, y). (f x, g y)) -` A \ space (M \\<^isub>M N))) \M)" + by (intro positive_integral_cong arg_cong2[where f=emeasure]) (auto simp: space_pair_measure) + also have "\ = emeasure (M \\<^isub>M N) ((\(x, y). (f x, g y)) -` A \ space (M \\<^isub>M N))" + using fg by (intro MN.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A]) + (auto cong: measurable_cong') + also have "\ = emeasure ?D A" + using fg A by (subst emeasure_distr) auto + finally show "emeasure ?P A = emeasure ?D A" . qed end \ No newline at end of file diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Mon Apr 23 12:14:35 2012 +0200 @@ -11,16 +11,14 @@ section "Generic Borel spaces" -definition "borel = sigma \ space = UNIV::'a::topological_space set, sets = {S. open S}\" -abbreviation "borel_measurable M \ measurable M borel" +definition borel :: "'a::topological_space measure" where + "borel = sigma UNIV {S. open S}" -interpretation borel: sigma_algebra borel - by (auto simp: borel_def intro!: sigma_algebra_sigma) +abbreviation "borel_measurable M \ measurable M borel" lemma in_borel_measurable: "f \ borel_measurable M \ - (\S \ sets (sigma \ space = UNIV, sets = {S. open S}\). - f -` S \ space M \ sets M)" + (\S \ sigma_sets UNIV {S. open S}. f -` S \ space M \ sets M)" by (auto simp add: measurable_def borel_def) lemma in_borel_measurable_borel: @@ -36,7 +34,7 @@ assumes "open A" shows "A \ sets borel" proof - have "A \ {S. open S}" unfolding mem_Collect_eq using assms . - thus ?thesis unfolding borel_def sigma_def by (auto intro!: sigma_sets.Basic) + thus ?thesis unfolding borel_def by auto qed lemma borel_closed[simp]: @@ -48,9 +46,9 @@ qed lemma borel_comp[intro,simp]: "A \ sets borel \ - A \ sets borel" - unfolding Compl_eq_Diff_UNIV by (intro borel.Diff) auto + unfolding Compl_eq_Diff_UNIV by (intro Diff) auto -lemma (in sigma_algebra) borel_measurable_vimage: +lemma borel_measurable_vimage: fixes f :: "'a \ 'x::t2_space" assumes borel: "f \ borel_measurable M" shows "f -` {x} \ space M \ sets M" @@ -65,12 +63,12 @@ thus ?thesis by auto qed -lemma (in sigma_algebra) borel_measurableI: +lemma borel_measurableI: fixes f :: "'a \ 'x\topological_space" assumes "\S. open S \ f -` S \ space M \ sets M" shows "f \ borel_measurable M" unfolding borel_def -proof (rule measurable_sigma, simp_all) +proof (rule measurable_measure_of, simp_all) fix S :: "'x set" assume "open S" thus "f -` S \ space M \ sets M" using assms[of S] by simp qed @@ -78,22 +76,22 @@ lemma borel_singleton[simp, intro]: fixes x :: "'a::t1_space" shows "A \ sets borel \ insert x A \ sets borel" - proof (rule borel.insert_in_sets) + proof (rule insert_in_sets) show "{x} \ sets borel" using closed_singleton[of x] by (rule borel_closed) qed simp -lemma (in sigma_algebra) borel_measurable_const[simp, intro]: +lemma borel_measurable_const[simp, intro]: "(\x. c) \ borel_measurable M" - by (auto intro!: measurable_const) + by auto -lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]: +lemma borel_measurable_indicator[simp, intro!]: assumes A: "A \ sets M" shows "indicator A \ borel_measurable M" unfolding indicator_def [abs_def] using A - by (auto intro!: measurable_If_set borel_measurable_const) + by (auto intro!: measurable_If_set) -lemma (in sigma_algebra) borel_measurable_indicator_iff: +lemma borel_measurable_indicator_iff: "(indicator A :: 'a \ 'x::{t1_space, zero_neq_one}) \ borel_measurable M \ A \ space M \ sets M" (is "?I \ borel_measurable M \ _") proof @@ -111,49 +109,7 @@ ultimately show "?I \ borel_measurable M" by auto qed -lemma (in sigma_algebra) borel_measurable_restricted: - fixes f :: "'a \ ereal" assumes "A \ sets M" - shows "f \ borel_measurable (restricted_space A) \ - (\x. f x * indicator A x) \ borel_measurable M" - (is "f \ borel_measurable ?R \ ?f \ borel_measurable M") -proof - - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) - have *: "f \ borel_measurable ?R \ ?f \ borel_measurable ?R" - by (auto intro!: measurable_cong) - show ?thesis unfolding * - unfolding in_borel_measurable_borel - proof (simp, safe) - fix S :: "ereal set" assume "S \ sets borel" - "\S\sets borel. ?f -` S \ A \ op \ A ` sets M" - then have "?f -` S \ A \ op \ A ` sets M" by auto - then have f: "?f -` S \ A \ sets M" - using `A \ sets M` sets_into_space by fastforce - show "?f -` S \ space M \ sets M" - proof cases - assume "0 \ S" - then have "?f -` S \ space M = ?f -` S \ A \ (space M - A)" - using `A \ sets M` sets_into_space by auto - then show ?thesis using f `A \ sets M` by (auto intro!: Un Diff) - next - assume "0 \ S" - then have "?f -` S \ space M = ?f -` S \ A" - using `A \ sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show ?thesis using f by auto - qed - next - fix S :: "ereal set" assume "S \ sets borel" - "\S\sets borel. ?f -` S \ space M \ sets M" - then have f: "?f -` S \ space M \ sets M" by auto - then show "?f -` S \ A \ op \ A ` sets M" - using `A \ sets M` sets_into_space - apply (simp add: image_iff) - apply (rule bexI[OF _ f]) - by auto - qed -qed - -lemma (in sigma_algebra) borel_measurable_subalgebra: +lemma borel_measurable_subalgebra: assumes "sets N \ sets M" "space N = space M" "f \ borel_measurable N" shows "f \ borel_measurable M" using assms unfolding measurable_def by auto @@ -220,7 +176,7 @@ shows "{x::'a::euclidean_space. x $$ i \ a} \ sets borel" by (auto intro!: borel_closed closed_halfspace_component_le) -lemma (in sigma_algebra) borel_measurable_less[simp, intro]: +lemma borel_measurable_less[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -233,7 +189,7 @@ by simp (blast intro: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_le[simp, intro]: +lemma borel_measurable_le[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -245,7 +201,7 @@ by simp blast qed -lemma (in sigma_algebra) borel_measurable_eq[simp, intro]: +lemma borel_measurable_eq[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -257,7 +213,7 @@ thus ?thesis using f g by auto qed -lemma (in sigma_algebra) borel_measurable_neq[simp, intro]: +lemma borel_measurable_neq[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -351,23 +307,70 @@ thus "x \ UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto qed auto -lemma halfspace_span_open: - "sigma_sets UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a})) - \ sets borel" - by (auto intro!: borel.sigma_sets_subset[simplified] borel_open - open_halfspace_component_lt) +lemma borel_sigma_sets_subset: + "A \ sets borel \ sigma_sets UNIV A \ sets borel" + using sigma_sets_subset[of A borel] by simp + +lemma borel_eq_sigmaI1: + fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\x. x \ X \ x \ sets (sigma UNIV (range F))" + assumes F: "\i. F i \ sets borel" + shows "borel = sigma UNIV (range F)" + unfolding borel_def +proof (intro sigma_eqI antisym) + have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" + unfolding borel_def by simp + also have "\ = sigma_sets UNIV X" + unfolding borel_eq by simp + also have "\ \ sigma_sets UNIV (range F)" + using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto + finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (range F)" . + show "sigma_sets UNIV (range F) \ sigma_sets UNIV {S. open S}" + unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto +qed auto -lemma halfspace_lt_in_halfspace: - "{x\'a. x $$ i < a} \ sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a})\)" - by (auto intro!: sigma_sets.Basic simp: sets_sigma) +lemma borel_eq_sigmaI2: + fixes F :: "'i \ 'j \ 'a::topological_space set" + and G :: "'l \ 'k \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV (range (\(i, j). G i j))" + assumes X: "\i j. G i j \ sets (sigma UNIV (range (\(i, j). F i j)))" + assumes F: "\i j. F i j \ sets borel" + shows "borel = sigma UNIV (range (\(i, j). F i j))" + using assms by (intro borel_eq_sigmaI1[where X="range (\(i, j). G i j)" and F="(\(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI3: + fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\x. x \ X \ x \ sets (sigma UNIV (range (\(i, j). F i j)))" + assumes F: "\i j. F i j \ sets borel" + shows "borel = sigma UNIV (range (\(i, j). F i j))" + using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI4: + fixes F :: "'i \ 'a::topological_space set" + and G :: "'l \ 'k \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV (range (\(i, j). G i j))" + assumes X: "\i j. G i j \ sets (sigma UNIV (range F))" + assumes F: "\i. F i \ sets borel" + shows "borel = sigma UNIV (range F)" + using assms by (intro borel_eq_sigmaI1[where X="range (\(i, j). G i j)" and F=F]) auto + +lemma borel_eq_sigmaI5: + fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV (range G)" + assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" + assumes F: "\i j. F i j \ sets borel" + shows "borel = sigma UNIV (range (\(i, j). F i j))" + using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto lemma halfspace_gt_in_halfspace: - "{x\'a. a < x $$ i} \ sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a})\)" - (is "?set \ sets ?SIGMA") + "{x\'a. a < x $$ i} \ sigma_sets UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a}))" + (is "?set \ ?SIGMA") proof - - interpret sigma_algebra "?SIGMA" - by (intro sigma_algebra_sigma_sets) (simp_all add: sets_sigma) - have *: "?set = (\n. space ?SIGMA - {x\'a. x $$ i < a + 1 / real (Suc n)})" + interpret sigma_algebra UNIV ?SIGMA + by (intro sigma_algebra_sigma_sets) simp_all + have *: "?set = (\n. UNIV - {x\'a. x $$ i < a + 1 / real (Suc n)})" proof (safe, simp_all add: not_less) fix x assume "a < x $$ i" with reals_Archimedean[of "x $$ i - a"] @@ -381,100 +384,78 @@ also assume "\ \ x" finally show "a < x" . qed - show "?set \ sets ?SIGMA" unfolding * - by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace) -qed - -lemma open_span_halfspace: - "sets borel \ sets (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp - then interpret sigma_algebra ?SIGMA . - { fix S :: "'a set" assume "S \ {S. open S}" - then have "open S" unfolding mem_Collect_eq . - from open_UNION[OF this] - obtain I where *: "S = - (\(a, b)\I. - (\ i op ! a)::'a) $$ i < x $$ i}) \ - (\ i op ! b)::'a) $$ i}))" - unfolding greaterThanLessThan_def - unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] - unfolding eucl_lessThan_eq_halfspaces[where 'a='a] - by blast - have "S \ sets ?SIGMA" - unfolding * - by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace) } - then show ?thesis unfolding borel_def - by (intro sets_sigma_subset) auto + show "?set \ ?SIGMA" unfolding * + by (auto intro!: Diff) qed -lemma halfspace_span_halfspace_le: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a})\) \ - sets (sigma \space=UNIV, sets=range (\ (a, i). {x. x $$ i \ a})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - { fix a i - have *: "{x::'a. x$$i < a} = (\n. {x. x$$i \ a - 1/real (Suc n)})" - proof (safe, simp_all) - fix x::'a assume *: "x$$i < a" - with reals_Archimedean[of "a - x$$i"] - obtain n where "x $$ i < a - 1 / (real (Suc n))" - by (auto simp: field_simps inverse_eq_divide) - then show "\n. x $$ i \ a - 1 / (real (Suc n))" - by (blast intro: less_imp_le) - next - fix x::'a and n - assume "x$$i \ a - 1 / real (Suc n)" - also have "\ < a" by auto - finally show "x$$i < a" . - qed - have "{x. x$$i < a} \ sets ?SIGMA" unfolding * - by (safe intro!: countable_UN) - (auto simp: sets_sigma intro!: sigma_sets.Basic) } - then show ?thesis by (intro sets_sigma_subset) auto -qed +lemma borel_eq_halfspace_less: + "borel = sigma UNIV (range (\(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI3[OF borel_def]) + fix S :: "'a set" assume "S \ {S. open S}" + then have "open S" by simp + from open_UNION[OF this] + obtain I where *: "S = + (\(a, b)\I. + (\ i op ! a)::'a) $$ i < x $$ i}) \ + (\ i op ! b)::'a) $$ i}))" + unfolding greaterThanLessThan_def + unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] + unfolding eucl_lessThan_eq_halfspaces[where 'a='a] + by blast + show "S \ ?SIGMA" + unfolding * + by (safe intro!: countable_UN Int countable_INT) (auto intro!: halfspace_gt_in_halfspace) +qed auto -lemma halfspace_span_halfspace_ge: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a})\) \ - sets (sigma \space=UNIV, sets=range (\ (a, i). {x. a \ x $$ i})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - { fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \ x$$i}" by auto - have "{x. x$$i < a} \ sets ?SIGMA" unfolding * - by (safe intro!: Diff) - (auto simp: sets_sigma intro!: sigma_sets.Basic) } - then show ?thesis by (intro sets_sigma_subset) auto -qed +lemma borel_eq_halfspace_le: + "borel = sigma UNIV (range (\ (a, i). {x::'a::ordered_euclidean_space. x $$ i \ a}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a i + have *: "{x::'a. x$$i < a} = (\n. {x. x$$i \ a - 1/real (Suc n)})" + proof (safe, simp_all) + fix x::'a assume *: "x$$i < a" + with reals_Archimedean[of "a - x$$i"] + obtain n where "x $$ i < a - 1 / (real (Suc n))" + by (auto simp: field_simps inverse_eq_divide) + then show "\n. x $$ i \ a - 1 / (real (Suc n))" + by (blast intro: less_imp_le) + next + fix x::'a and n + assume "x$$i \ a - 1 / real (Suc n)" + also have "\ < a" by auto + finally show "x$$i < a" . + qed + show "{x. x$$i < a} \ ?SIGMA" unfolding * + by (safe intro!: countable_UN) auto +qed auto -lemma halfspace_le_span_halfspace_gt: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i \ a})\) \ - sets (sigma \space=UNIV, sets=range (\ (a, i). {x. a < x $$ i})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - { fix a i have *: "{x::'a. x$$i \ a} = space ?SIGMA - {x::'a. a < x$$i}" by auto - have "{x. x$$i \ a} \ sets ?SIGMA" unfolding * - by (safe intro!: Diff) - (auto simp: sets_sigma intro!: sigma_sets.Basic) } - then show ?thesis by (intro sets_sigma_subset) auto -qed +lemma borel_eq_halfspace_ge: + "borel = sigma UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. a \ x $$ i}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \ x$$i}" by auto + show "{x. x$$i < a} \ ?SIGMA" unfolding * + by (safe intro!: compl_sets) auto +qed auto -lemma halfspace_le_span_atMost: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i \ a})\) \ - sets (sigma \space=UNIV, sets=range (\a. {..a\'a\ordered_euclidean_space})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - have "\a i. {x. x$$i \ a} \ sets ?SIGMA" +lemma borel_eq_halfspace_greater: + "borel = sigma UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. a < x $$ i}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) + fix a i have *: "{x::'a. x$$i \ a} = space ?SIGMA - {x::'a. a < x$$i}" by auto + show "{x. x$$i \ a} \ ?SIGMA" unfolding * + by (safe intro!: compl_sets) auto +qed auto + +lemma borel_eq_atMost: + "borel = sigma UNIV (range (\a. {..a\'a\ordered_euclidean_space}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a i show "{x. x$$i \ a} \ ?SIGMA" proof cases - fix a i assume "i < DIM('a)" + assume "i < DIM('a)" then have *: "{x::'a. x$$i \ a} = (\k::nat. {.. (\\ n. if n = i then a else real k)})" proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) fix x @@ -484,28 +465,19 @@ then show "\k::nat. \ia. ia \ i \ ia < DIM('a) \ x $$ ia \ real k" by (auto intro!: exI[of _ k]) qed - show "{x. x$$i \ a} \ sets ?SIGMA" unfolding * - by (safe intro!: countable_UN) - (auto simp: sets_sigma intro!: sigma_sets.Basic) - next - fix a i assume "\ i < DIM('a)" - then show "{x. x$$i \ a} \ sets ?SIGMA" - using top by auto - qed - then show ?thesis by (intro sets_sigma_subset) auto -qed + show "{x. x$$i \ a} \ ?SIGMA" unfolding * + by (safe intro!: countable_UN) auto + qed (auto intro: sigma_sets_top sigma_sets.Empty) +qed auto -lemma halfspace_le_span_greaterThan: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i \ a})\) \ - sets (sigma \space=UNIV, sets=range (\a. {a<..})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - have "\a i. {x. x$$i \ a} \ sets ?SIGMA" +lemma borel_eq_greaterThan: + "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {a<..}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a i show "{x. x$$i \ a} \ ?SIGMA" proof cases - fix a i assume "i < DIM('a)" - have "{x::'a. x$$i \ a} = space ?SIGMA - {x::'a. a < x$$i}" by auto + assume "i < DIM('a)" + have "{x::'a. x$$i \ a} = UNIV - {x::'a. a < x$$i}" by auto also have *: "{x::'a. a < x$$i} = (\k::nat. {(\\ n. if n = i then a else -real k) <..})" using `i -real k < x $$ ia" by (auto intro!: exI[of _ k]) qed - finally show "{x. x$$i \ a} \ sets ?SIGMA" + finally show "{x. x$$i \ a} \ ?SIGMA" apply (simp only:) apply (safe intro!: countable_UN Diff) - apply (auto simp: sets_sigma intro!: sigma_sets.Basic) + apply (auto intro: sigma_sets_top) done - next - fix a i assume "\ i < DIM('a)" - then show "{x. x$$i \ a} \ sets ?SIGMA" - using top by auto - qed - then show ?thesis by (intro sets_sigma_subset) auto -qed + qed (auto intro: sigma_sets_top sigma_sets.Empty) +qed auto -lemma halfspace_le_span_lessThan: - "sets (sigma \space=UNIV, sets=range (\ (a, i). {x\'a\ordered_euclidean_space. a \ x $$ i})\) \ - sets (sigma \space=UNIV, sets=range (\a. {..)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - have "\a i. {x. a \ x$$i} \ sets ?SIGMA" +lemma borel_eq_lessThan: + "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {.. x$$i} \ ?SIGMA" proof cases fix a i assume "i < DIM('a)" - have "{x::'a. a \ x$$i} = space ?SIGMA - {x::'a. x$$i < a}" by auto + have "{x::'a. a \ x$$i} = UNIV - {x::'a. x$$i < a}" by auto also have *: "{x::'a. x$$i < a} = (\k::nat. {..< (\\ n. if n = i then a else real k)})" using `i x $$ ia < real k" by (auto intro!: exI[of _ k]) qed - finally show "{x. a \ x$$i} \ sets ?SIGMA" + finally show "{x. a \ x$$i} \ ?SIGMA" apply (simp only:) apply (safe intro!: countable_UN Diff) - apply (auto simp: sets_sigma intro!: sigma_sets.Basic) + apply (auto intro: sigma_sets_top) done - next - fix a i assume "\ i < DIM('a)" - then show "{x. a \ x$$i} \ sets ?SIGMA" - using top by auto - qed - then show ?thesis by (intro sets_sigma_subset) auto -qed - -lemma atMost_span_atLeastAtMost: - "sets (sigma \space=UNIV, sets=range (\a. {..a\'a\ordered_euclidean_space})\) \ - sets (sigma \space=UNIV, sets=range (\(a,b). {a..b})\)" - (is "_ \ sets ?SIGMA") -proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - { fix a::'a - have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" - proof (safe, simp_all add: eucl_le[where 'a='a]) - fix x - from real_arch_simple[of "Max ((\i. - x$$i)`{.. real k" - by (subst (asm) Max_le_iff) (auto simp: field_simps) - then have "- real k \ x$$i" by simp } - then show "\n::nat. \i x $$ i" - by (auto intro!: exI[of _ k]) - qed - have "{..a} \ sets ?SIGMA" unfolding * - by (safe intro!: countable_UN) - (auto simp: sets_sigma intro!: sigma_sets.Basic) } - then show ?thesis by (intro sets_sigma_subset) auto -qed - -lemma borel_eq_atMost: - "borel = (sigma \space=UNIV, sets=range (\ a. {.. a::'a\ordered_euclidean_space})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace - by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto + qed (auto intro: sigma_sets_top sigma_sets.Empty) qed auto lemma borel_eq_atLeastAtMost: - "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space, b). {a .. b})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using atMost_span_atLeastAtMost halfspace_le_span_atMost - halfspace_span_halfspace_le open_span_halfspace - by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto -qed auto - -lemma borel_eq_greaterThan: - "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space). {a <..})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using halfspace_le_span_greaterThan - halfspace_span_halfspace_le open_span_halfspace - by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto -qed auto - -lemma borel_eq_lessThan: - "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space). {..< a})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using halfspace_le_span_lessThan - halfspace_span_halfspace_ge open_span_halfspace - by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto + "borel = sigma UNIV (range (\(a,b). {a..b} \'a\ordered_euclidean_space set))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) + fix a::'a + have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" + proof (safe, simp_all add: eucl_le[where 'a='a]) + fix x + from real_arch_simple[of "Max ((\i. - x$$i)`{.. real k" + by (subst (asm) Max_le_iff) (auto simp: field_simps) + then have "- real k \ x$$i" by simp } + then show "\n::nat. \i x $$ i" + by (auto intro!: exI[of _ k]) + qed + show "{..a} \ ?SIGMA" unfolding * + by (safe intro!: countable_UN) + (auto intro!: sigma_sets_top) qed auto lemma borel_eq_greaterThanLessThan: - "borel = (sigma \space=UNIV, sets=range (\ (a, b). {a <..< (b :: 'a \ ordered_euclidean_space)})\)" + "borel = sigma UNIV (range (\ (a, b). {a <..< b} :: 'a \ ordered_euclidean_space set))" (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto - show "sets borel \ sets ?SIGMA" - proof - - have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto - then interpret sigma_algebra ?SIGMA . - { fix M :: "'a set" assume "M \ {S. open S}" - then have "open M" by simp - have "M \ sets ?SIGMA" - apply (subst open_UNION[OF `open M`]) - apply (safe intro!: countable_UN) - apply (auto simp add: sigma_def intro!: sigma_sets.Basic) - done } - then show ?thesis - unfolding borel_def by (intro sets_sigma_subset) auto - qed +proof (rule borel_eq_sigmaI1[OF borel_def]) + fix M :: "'a set" assume "M \ {S. open S}" + then have "open M" by simp + show "M \ ?SIGMA" + apply (subst open_UNION[OF `open M`]) + apply (safe intro!: countable_UN) + apply auto + done qed auto lemma borel_eq_atLeastLessThan: - "borel = sigma \space=UNIV, sets=range (\(a, b). {a ..< b :: real})\" (is "_ = ?S") -proof (intro algebra.equality antisym) - interpret sigma_algebra ?S - by (rule sigma_algebra_sigma) auto - show "sets borel \ sets ?S" - unfolding borel_eq_lessThan - proof (intro sets_sigma_subset subsetI) - have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto - fix A :: "real set" assume "A \ sets \space = UNIV, sets = range lessThan\" - then obtain x where "A = {..< x}" by auto - then have "A = (\i::nat. {-real i ..< x})" - by (auto simp: move_uminus real_arch_simple) - then show "A \ sets ?S" - by (auto simp: sets_sigma intro!: sigma_sets.intros) - qed simp - show "sets ?S \ sets borel" - by (intro borel.sets_sigma_subset) auto -qed simp_all - -lemma borel_eq_halfspace_le: - "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. x$$i \ a})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using open_span_halfspace halfspace_span_halfspace_le by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto + "borel = sigma UNIV (range (\(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) + have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto + fix x :: real + have "{..i::nat. {-real i ..< x})" + by (auto simp: move_uminus real_arch_simple) + then show "{..< x} \ ?SIGMA" + by (auto intro: sigma_sets.intros) qed auto -lemma borel_eq_halfspace_less: - "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. x$$i < a})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using open_span_halfspace . - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto -qed auto - -lemma borel_eq_halfspace_gt: - "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. a < x$$i})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto -qed auto - -lemma borel_eq_halfspace_ge: - "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. a \ x$$i})\)" - (is "_ = ?SIGMA") -proof (intro algebra.equality antisym) - show "sets borel \ sets ?SIGMA" - using halfspace_span_halfspace_ge open_span_halfspace by auto - show "sets ?SIGMA \ sets borel" - by (rule borel.sets_sigma_subset) auto -qed auto - -lemma (in sigma_algebra) borel_measurable_halfspacesI: +lemma borel_measurable_halfspacesI: fixes f :: "'a \ 'c\ordered_euclidean_space" - assumes "borel = (sigma \space=UNIV, sets=range F\)" - and "\a i. S a i = f -` F (a,i) \ space M" - and "\a i. \ i < DIM('c) \ S a i \ sets M" + assumes F: "borel = sigma UNIV (range F)" + and S_eq: "\a i. S a i = f -` F (a,i) \ space M" + and S: "\a i. \ i < DIM('c) \ S a i \ sets M" shows "f \ borel_measurable M = (\ia::real. S a i \ sets M)" proof safe fix a :: real and i assume i: "i < DIM('c)" and f: "f \ borel_measurable M" then show "S a i \ sets M" unfolding assms - by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def) + by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1)) next assume a: "\ia. S a i \ sets M" { fix a i have "S a i \ sets M" @@ -740,61 +590,58 @@ with a show ?thesis unfolding assms(2) by simp next assume "\ i < DIM('c)" - from assms(3)[OF this] show ?thesis . + from S[OF this] show ?thesis . qed } - then have "f \ measurable M (sigma \space=UNIV, sets=range F\)" - by (auto intro!: measurable_sigma simp: assms(2)) - then show "f \ borel_measurable M" unfolding measurable_def - unfolding assms(1) by simp + then show "f \ borel_measurable M" + by (auto intro!: measurable_measure_of simp: S_eq F) qed -lemma (in sigma_algebra) borel_measurable_iff_halfspace_le: +lemma borel_measurable_iff_halfspace_le: fixes f :: "'a \ 'c\ordered_euclidean_space" shows "f \ borel_measurable M = (\ia. {w \ space M. f w $$ i \ a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto -lemma (in sigma_algebra) borel_measurable_iff_halfspace_less: +lemma borel_measurable_iff_halfspace_less: fixes f :: "'a \ 'c\ordered_euclidean_space" shows "f \ borel_measurable M \ (\ia. {w \ space M. f w $$ i < a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto -lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge: +lemma borel_measurable_iff_halfspace_ge: fixes f :: "'a \ 'c\ordered_euclidean_space" shows "f \ borel_measurable M = (\ia. {w \ space M. a \ f w $$ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto -lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater: +lemma borel_measurable_iff_halfspace_greater: fixes f :: "'a \ 'c\ordered_euclidean_space" shows "f \ borel_measurable M \ (\ia. {w \ space M. a < f w $$ i} \ sets M)" - by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_gt]) auto + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto -lemma (in sigma_algebra) borel_measurable_iff_le: +lemma borel_measurable_iff_le: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" using borel_measurable_iff_halfspace_le[where 'c=real] by simp -lemma (in sigma_algebra) borel_measurable_iff_less: +lemma borel_measurable_iff_less: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" using borel_measurable_iff_halfspace_less[where 'c=real] by simp -lemma (in sigma_algebra) borel_measurable_iff_ge: +lemma borel_measurable_iff_ge: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" using borel_measurable_iff_halfspace_ge[where 'c=real] by simp -lemma (in sigma_algebra) borel_measurable_iff_greater: +lemma borel_measurable_iff_greater: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" using borel_measurable_iff_halfspace_greater[where 'c=real] by simp lemma borel_measurable_euclidean_component: "(\x::'a::euclidean_space. x $$ i) \ borel_measurable borel" - unfolding borel_def[where 'a=real] -proof (rule borel.measurable_sigma, simp_all) +proof (rule borel_measurableI) fix S::"real set" assume "open S" from open_vimage_euclidean_component[OF this] - show "(\x. x $$ i) -` S \ sets borel" + show "(\x. x $$ i) -` S \ space borel \ sets borel" by (auto intro: borel_open) qed -lemma (in sigma_algebra) borel_measurable_euclidean_space: +lemma borel_measurable_euclidean_space: fixes f :: "'a \ 'c::ordered_euclidean_space" shows "f \ borel_measurable M \ (\ix. f x $$ i) \ borel_measurable M)" proof safe @@ -810,7 +657,7 @@ subsection "Borel measurable operators" -lemma (in sigma_algebra) affine_borel_measurable_vector: +lemma affine_borel_measurable_vector: fixes f :: "'a \ 'x::real_normed_vector" assumes "f \ borel_measurable M" shows "(\x. a + b *\<^sub>R f x) \ borel_measurable M" @@ -821,8 +668,7 @@ assume "b \ 0" with `open S` have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") by (auto intro!: open_affinity simp: scaleR_add_right) - hence "?S \ sets borel" - unfolding borel_def by (auto simp: sigma_def intro!: sigma_sets.Basic) + hence "?S \ sets borel" by auto moreover from `b \ 0` have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) @@ -831,13 +677,13 @@ qed simp qed -lemma (in sigma_algebra) affine_borel_measurable: +lemma affine_borel_measurable: fixes g :: "'a \ real" assumes g: "g \ borel_measurable M" shows "(\x. a + (g x) * b) \ borel_measurable M" using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute) -lemma (in sigma_algebra) borel_measurable_add[simp, intro]: +lemma borel_measurable_add[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -855,7 +701,7 @@ by (simp add: borel_measurable_iff_ge) qed -lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]: +lemma borel_measurable_setsum[simp, intro]: fixes f :: "'c \ 'a \ real" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" @@ -864,7 +710,7 @@ thus ?thesis using assms by induct auto qed simp -lemma (in sigma_algebra) borel_measurable_square: +lemma borel_measurable_square: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" shows "(\x. (f x)^2) \ borel_measurable M" @@ -916,7 +762,7 @@ shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4" by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric]) -lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]: +lemma borel_measurable_uminus[simp, intro]: fixes g :: "'a \ real" assumes g: "g \ borel_measurable M" shows "(\x. - g x) \ borel_measurable M" @@ -928,7 +774,7 @@ finally show ?thesis . qed -lemma (in sigma_algebra) borel_measurable_times[simp, intro]: +lemma borel_measurable_times[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -947,7 +793,7 @@ using 1 2 by simp qed -lemma (in sigma_algebra) borel_measurable_setprod[simp, intro]: +lemma borel_measurable_setprod[simp, intro]: fixes f :: "'c \ 'a \ real" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" @@ -956,14 +802,14 @@ thus ?thesis using assms by induct auto qed simp -lemma (in sigma_algebra) borel_measurable_diff[simp, intro]: +lemma borel_measurable_diff[simp, intro]: fixes f :: "'a \ real" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x - g x) \ borel_measurable M" unfolding diff_minus using assms by fast -lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]: +lemma borel_measurable_inverse[simp, intro]: fixes f :: "'a \ real" assumes "f \ borel_measurable M" shows "(\x. inverse (f x)) \ borel_measurable M" @@ -978,7 +824,7 @@ by (auto intro!: Int Un) qed -lemma (in sigma_algebra) borel_measurable_divide[simp, intro]: +lemma borel_measurable_divide[simp, intro]: fixes f :: "'a \ real" assumes "f \ borel_measurable M" and "g \ borel_measurable M" @@ -986,7 +832,7 @@ unfolding field_divide_inverse by (rule borel_measurable_inverse borel_measurable_times assms)+ -lemma (in sigma_algebra) borel_measurable_max[intro, simp]: +lemma borel_measurable_max[intro, simp]: fixes f g :: "'a \ real" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" @@ -1001,7 +847,7 @@ by (auto intro!: Int) qed -lemma (in sigma_algebra) borel_measurable_min[intro, simp]: +lemma borel_measurable_min[intro, simp]: fixes f g :: "'a \ real" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" @@ -1016,7 +862,7 @@ by (auto intro!: Int) qed -lemma (in sigma_algebra) borel_measurable_abs[simp, intro]: +lemma borel_measurable_abs[simp, intro]: assumes "f \ borel_measurable M" shows "(\x. \f x :: real\) \ borel_measurable M" proof - @@ -1033,14 +879,14 @@ fixes f :: "'a::topological_space \ 'b::t1_space" assumes "continuous_on UNIV f" shows "f \ borel_measurable borel" - apply(rule borel.borel_measurableI) + apply(rule borel_measurableI) using continuous_open_preimage[OF assms] unfolding vimage_def by auto lemma borel_measurable_continuous_on: fixes f :: "'a::topological_space \ 'b::t1_space" assumes cont: "continuous_on A f" "open A" shows "(\x. if x \ A then f x else c) \ borel_measurable borel" (is "?f \ _") -proof (rule borel.borel_measurableI) +proof (rule borel_measurableI) fix S :: "'b set" assume "open S" then have "open {x\A. f x \ S}" by (intro continuous_open_preimage[OF cont]) auto @@ -1049,11 +895,11 @@ {x\A. f x \ S} \ (if c \ S then space borel - A else {})" by (auto split: split_if_asm) also have "\ \ sets borel" - using * `open A` by (auto simp del: space_borel intro!: borel.Un) + using * `open A` by (auto simp del: space_borel intro!: Un) finally show "?f -` S \ space borel \ sets borel" . qed -lemma (in sigma_algebra) convex_measurable: +lemma convex_measurable: fixes a b :: real assumes X: "X \ borel_measurable M" "X ` space M \ { a <..< b}" assumes q: "convex_on { a <..< b} q" @@ -1091,7 +937,7 @@ finally show ?thesis . qed -lemma (in sigma_algebra) borel_measurable_log[simp,intro]: +lemma borel_measurable_log[simp,intro]: assumes f: "f \ borel_measurable M" and "1 < b" shows "(\x. log b (f x)) \ borel_measurable M" using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]] @@ -1101,25 +947,21 @@ lemma borel_measurable_ereal_borel: "ereal \ borel_measurable borel" - unfolding borel_def[where 'a=ereal] -proof (rule borel.measurable_sigma) - fix X :: "ereal set" assume "X \ sets \space = UNIV, sets = {S. open S} \" - then have "open X" by simp +proof (rule borel_measurableI) + fix X :: "ereal set" assume "open X" then have "open (ereal -` X \ space borel)" by (simp add: open_ereal_vimage) then show "ereal -` X \ space borel \ sets borel" by auto -qed auto +qed -lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]: +lemma borel_measurable_ereal[simp, intro]: assumes f: "f \ borel_measurable M" shows "(\x. ereal (f x)) \ borel_measurable M" using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def . lemma borel_measurable_real_of_ereal_borel: "(real :: ereal \ real) \ borel_measurable borel" - unfolding borel_def[where 'a=real] -proof (rule borel.measurable_sigma) - fix B :: "real set" assume "B \ sets \space = UNIV, sets = {S. open S} \" - then have "open B" by simp +proof (rule borel_measurableI) + fix B :: "real set" assume "open B" have *: "ereal -` real -` (B - {0}) = B - {0}" by auto have open_real: "open (real -` (B - {0}) :: ereal set)" unfolding open_ereal_def * using `open B` by auto @@ -1137,13 +979,13 @@ then show "(real -` B :: ereal set) \ space borel \ sets borel" using open_real by auto qed -qed auto +qed -lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]: +lemma borel_measurable_real_of_ereal[simp, intro]: assumes f: "f \ borel_measurable M" shows "(\x. real (f x :: ereal)) \ borel_measurable M" using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def . -lemma (in sigma_algebra) borel_measurable_ereal_iff: +lemma borel_measurable_ereal_iff: shows "(\x. ereal (f x)) \ borel_measurable M \ f \ borel_measurable M" proof assume "(\x. ereal (f x)) \ borel_measurable M" @@ -1151,7 +993,7 @@ show "f \ borel_measurable M" by auto qed auto -lemma (in sigma_algebra) borel_measurable_ereal_iff_real: +lemma borel_measurable_ereal_iff_real: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ ((\x. real (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" @@ -1165,7 +1007,7 @@ finally show "f \ borel_measurable M" . qed (auto intro: measurable_sets borel_measurable_real_of_ereal) -lemma (in sigma_algebra) less_eq_ge_measurable: +lemma less_eq_ge_measurable: fixes f :: "'a \ 'c::linorder" shows "f -` {a <..} \ space M \ sets M \ f -` {..a} \ space M \ sets M" proof @@ -1178,7 +1020,7 @@ ultimately show "f -` {a <..} \ space M \ sets M" by auto qed -lemma (in sigma_algebra) greater_eq_le_measurable: +lemma greater_eq_le_measurable: fixes f :: "'a \ 'c::linorder" shows "f -` {..< a} \ space M \ sets M \ f -` {a ..} \ space M \ sets M" proof @@ -1191,28 +1033,27 @@ ultimately show "f -` {a ..} \ space M \ sets M" by auto qed -lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal: +lemma borel_measurable_uminus_borel_ereal: "(uminus :: ereal \ ereal) \ borel_measurable borel" -proof (subst borel_def, rule borel.measurable_sigma) - fix X :: "ereal set" assume "X \ sets \space = UNIV, sets = {S. open S}\" - then have "open X" by simp +proof (rule borel_measurableI) + fix X :: "ereal set" assume "open X" have "uminus -` X = uminus ` X" by (force simp: image_iff) then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto then show "uminus -` X \ space borel \ sets borel" by auto -qed auto +qed -lemma (in sigma_algebra) borel_measurable_uminus_ereal[intro]: +lemma borel_measurable_uminus_ereal[intro]: assumes "f \ borel_measurable M" shows "(\x. - f x :: ereal) \ borel_measurable M" using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def) -lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]: +lemma borel_measurable_uminus_eq_ereal[simp]: "(\x. - f x :: ereal) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp qed auto -lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal: +lemma borel_measurable_eq_atMost_ereal: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" proof (intro iffI allI) @@ -1244,7 +1085,7 @@ qed qed (simp add: measurable_sets) -lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal: +lemma borel_measurable_eq_atLeast_ereal: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" proof assume pos: "\a. f -` {a..} \ space M \ sets M" @@ -1255,15 +1096,15 @@ then show "f \ borel_measurable M" by simp qed (simp add: measurable_sets) -lemma (in sigma_algebra) borel_measurable_ereal_iff_less: +lemma borel_measurable_ereal_iff_less: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable .. -lemma (in sigma_algebra) borel_measurable_ereal_iff_ge: +lemma borel_measurable_ereal_iff_ge: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable .. -lemma (in sigma_algebra) borel_measurable_ereal_eq_const: +lemma borel_measurable_ereal_eq_const: fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" shows "{x\space M. f x = c} \ sets M" proof - @@ -1271,7 +1112,7 @@ then show ?thesis using assms by (auto intro!: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_ereal_neq_const: +lemma borel_measurable_ereal_neq_const: fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" shows "{x\space M. f x \ c} \ sets M" proof - @@ -1279,7 +1120,7 @@ then show ?thesis using assms by (auto intro!: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]: +lemma borel_measurable_ereal_le[intro,simp]: fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -1294,7 +1135,7 @@ with f g show ?thesis by (auto intro!: Un simp: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]: +lemma borel_measurable_ereal_less[intro,simp]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -1304,7 +1145,7 @@ then show ?thesis using g f by auto qed -lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]: +lemma borel_measurable_ereal_eq[intro,simp]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -1314,7 +1155,7 @@ then show ?thesis using g f by auto qed -lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]: +lemma borel_measurable_ereal_neq[intro,simp]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" @@ -1324,12 +1165,12 @@ thus ?thesis using f g by auto qed -lemma (in sigma_algebra) split_sets: +lemma split_sets: "{x\space M. P x \ Q x} = {x\space M. P x} \ {x\space M. Q x}" "{x\space M. P x \ Q x} = {x\space M. P x} \ {x\space M. Q x}" by auto -lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]: +lemma borel_measurable_ereal_add[intro, simp]: fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" "g \ borel_measurable M" shows "(\x. f x + g x) \ borel_measurable M" @@ -1344,7 +1185,7 @@ intro!: Un measurable_If measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]: +lemma borel_measurable_ereal_setsum[simp, intro]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" @@ -1354,7 +1195,7 @@ by induct auto qed (simp add: borel_measurable_const) -lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]: +lemma borel_measurable_ereal_abs[intro, simp]: fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" shows "(\x. \f x\) \ borel_measurable M" proof - @@ -1362,7 +1203,7 @@ then show ?thesis using assms by (auto intro!: measurable_If) qed -lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]: +lemma borel_measurable_ereal_times[intro, simp]: fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" "g \ borel_measurable M" shows "(\x. f x * g x) \ borel_measurable M" proof - @@ -1386,7 +1227,7 @@ (auto simp: split_sets intro!: Int) qed -lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]: +lemma borel_measurable_ereal_setprod[simp, intro]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" @@ -1395,21 +1236,21 @@ thus ?thesis using assms by induct auto qed simp -lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]: +lemma borel_measurable_ereal_min[simp, intro]: fixes f g :: "'a \ ereal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. min (g x) (f x)) \ borel_measurable M" using assms unfolding min_def by (auto intro!: measurable_If) -lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]: +lemma borel_measurable_ereal_max[simp, intro]: fixes f g :: "'a \ ereal" assumes "f \ borel_measurable M" and "g \ borel_measurable M" shows "(\x. max (g x) (f x)) \ borel_measurable M" using assms unfolding max_def by (auto intro!: measurable_If) -lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]: +lemma borel_measurable_SUP[simp, intro]: fixes f :: "'d\countable \ 'a \ ereal" assumes "\i. i \ A \ f i \ borel_measurable M" shows "(\x. SUP i : A. f i x) \ borel_measurable M" (is "?sup \ borel_measurable M") @@ -1422,7 +1263,7 @@ using assms by auto qed -lemma (in sigma_algebra) borel_measurable_INF[simp, intro]: +lemma borel_measurable_INF[simp, intro]: fixes f :: "'d :: countable \ 'a \ ereal" assumes "\i. i \ A \ f i \ borel_measurable M" shows "(\x. INF i : A. f i x) \ borel_measurable M" (is "?inf \ borel_measurable M") @@ -1435,26 +1276,39 @@ using assms by auto qed -lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]: +lemma borel_measurable_liminf[simp, intro]: fixes f :: "nat \ 'a \ ereal" assumes "\i. f i \ borel_measurable M" shows "(\x. liminf (\i. f i x)) \ borel_measurable M" unfolding liminf_SUPR_INFI using assms by auto -lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]: +lemma borel_measurable_limsup[simp, intro]: fixes f :: "nat \ 'a \ ereal" assumes "\i. f i \ borel_measurable M" shows "(\x. limsup (\i. f i x)) \ borel_measurable M" unfolding limsup_INFI_SUPR using assms by auto -lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]: +lemma borel_measurable_ereal_diff[simp, intro]: fixes f g :: "'a \ ereal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. f x - g x) \ borel_measurable M" unfolding minus_ereal_def using assms by auto -lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]: +lemma borel_measurable_ereal_inverse[simp, intro]: + assumes f: "f \ borel_measurable M" shows "(\x. inverse (f x) :: ereal) \ borel_measurable M" +proof - + { fix x have "inverse (f x) = (if f x = 0 then \ else ereal (inverse (real (f x))))" + by (cases "f x") auto } + with f show ?thesis + by (auto intro!: measurable_If) +qed + +lemma borel_measurable_ereal_divide[simp, intro]: + "f \ borel_measurable M \ g \ borel_measurable M \ (\x. f x / g x :: ereal) \ borel_measurable M" + unfolding divide_ereal_def by auto + +lemma borel_measurable_psuminf[simp, intro]: fixes f :: "nat \ 'a \ ereal" assumes "\i. f i \ borel_measurable M" and pos: "\i x. x \ space M \ 0 \ f i x" shows "(\x. (\i. f i x)) \ borel_measurable M" @@ -1465,7 +1319,7 @@ section "LIMSEQ is borel measurable" -lemma (in sigma_algebra) borel_measurable_LIMSEQ: +lemma borel_measurable_LIMSEQ: fixes u :: "nat \ 'a \ real" assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" and u: "\i. u i \ borel_measurable M" diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Caratheodory.thy Mon Apr 23 12:14:35 2012 +0200 @@ -6,7 +6,7 @@ header {*Caratheodory Extension Theorem*} theory Caratheodory -imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits" + imports Measure_Space begin lemma sums_def2: @@ -53,128 +53,53 @@ subsection {* Measure Spaces *} -record 'a measure_space = "'a algebra" + - measure :: "'a set \ ereal" - -definition positive where "positive M f \ f {} = (0::ereal) \ (\A\sets M. 0 \ f A)" - -definition additive where "additive M f \ - (\x \ sets M. \y \ sets M. x \ y = {} \ f (x \ y) = f x + f y)" - -definition countably_additive :: "('a, 'b) algebra_scheme \ ('a set \ ereal) \ bool" where - "countably_additive M f \ (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ - (\i. f (A i)) = f (\i. A i))" - -definition increasing where "increasing M f \ - (\x \ sets M. \y \ sets M. x \ y \ f x \ f y)" - definition subadditive where "subadditive M f \ - (\x \ sets M. \y \ sets M. x \ y = {} \ f (x \ y) \ f x + f y)" + (\x\M. \y\M. x \ y = {} \ f (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 \ + (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" -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 lambda_system where "lambda_system \ M f = {l \ M. + \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" 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) \ (\i. f (A i)) = r}" - -locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" + - assumes measure_positive: "positive M (measure M)" - and ca: "countably_additive M (measure M)" - -abbreviation (in measure_space) "\ \ measure M" - -lemma (in measure_space) - shows empty_measure[simp, intro]: "\ {} = 0" - and positive_measure[simp, intro!]: "\A. A \ sets M \ 0 \ \ A" - using measure_positive unfolding positive_def by auto - -lemma increasingD: - "increasing M f \ x \ y \ x\sets M \ y\sets M \ f x \ f y" - by (auto simp add: increasing_def) + \A. range A \ M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) = r}" 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 \ M \ y \ 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" - by (auto simp add: additive_def) - -lemma countably_additiveI: - assumes "\A x. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M - \ (\i. f (A i)) = f (\i. A i)" - shows "countably_additive M f" - using assms by (simp add: countably_additive_def) - -section "Extend binary sets" - -lemma LIMSEQ_binaryset: - assumes f: "f {} = 0" - shows "(\n. \i f A + f B" -proof - - have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" - proof - fix n - show "(\i < Suc (Suc n). f (binaryset A B i)) = f A + f B" - by (induct n) (auto simp add: binaryset_def f) - qed - moreover - have "... ----> f A + f B" by (rule tendsto_const) - ultimately - have "(\n. \i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B" - by metis - hence "(\n. \i< n+2. f (binaryset A B i)) ----> f A + f B" - by simp - thus ?thesis by (rule LIMSEQ_offset [where k=2]) -qed - -lemma binaryset_sums: - assumes f: "f {} = 0" - shows "(\n. f (binaryset A B n)) sums (f A + f B)" - by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) - -lemma suminf_binaryset_eq: - fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" - shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" - by (metis binaryset_sums sums_unique) - subsection {* Lambda Systems *} lemma (in algebra) lambda_system_eq: - shows "lambda_system M f = {l \ sets M. - \x \ sets M. f (x \ l) + f (x - l) = f x}" + shows "lambda_system \ M f = {l \ M. \x \ 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" + have [simp]: "!!l x. l \ M \ x \ M \ (\ - l) \ x = x - l" by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) show ?thesis by (auto simp add: lambda_system_def) (metis Int_commute)+ qed lemma (in algebra) lambda_system_empty: - "positive M f \ {} \ lambda_system M f" + "positive M f \ {} \ lambda_system \ M f" by (auto simp add: positive_def lambda_system_eq) lemma lambda_system_sets: - "x \ lambda_system M f \ x \ sets M" + "x \ lambda_system \ M f \ x \ M" by (simp add: lambda_system_def) lemma (in algebra) lambda_system_Compl: fixes f:: "'a set \ ereal" - assumes x: "x \ lambda_system M f" - shows "space M - x \ lambda_system M f" + assumes x: "x \ lambda_system \ M f" + shows "\ - x \ lambda_system \ M f" proof - - have "x \ space M" + have "x \ \" by (metis sets_into_space lambda_system_sets x) - hence "space M - (space M - x) = x" + hence "\ - (\ - x) = x" by (metis double_diff equalityE) with x show ?thesis by (force simp add: lambda_system_def ac_simps) @@ -182,16 +107,16 @@ lemma (in algebra) lambda_system_Int: fixes f:: "'a set \ ereal" - assumes xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" - shows "x \ y \ lambda_system M f" + 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" + assume x: "x \ M" and y: "y \ M" and u: "u \ M" + and fx: "\z\M. f (z \ x) + f (z - x) = f z" + and fy: "\z\M. f (z \ y) + f (z - y) = f z" + have "u - x \ y \ M" by (metis Diff Diff_Int Un u x y) moreover have "(u - (x \ y)) \ y = u \ y - x" by blast @@ -216,20 +141,20 @@ lemma (in algebra) lambda_system_Un: fixes f:: "'a set \ ereal" - assumes xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" - shows "x \ y \ lambda_system M f" + assumes xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" + shows "x \ y \ lambda_system \ M f" proof - - have "(space M - x) \ (space M - y) \ sets M" + have "(\ - x) \ (\ - y) \ M" by (metis Diff_Un Un compl_sets lambda_system_sets xl yl) moreover - have "x \ y = space M - ((space M - x) \ (space M - y))" + have "x \ y = \ - ((\ - x) \ (\ - y))" by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+ ultimately show ?thesis by (metis lambda_system_Compl lambda_system_Int xl yl) qed lemma (in algebra) lambda_system_algebra: - "positive M f \ algebra (M\sets := lambda_system M f\)" + "positive M f \ algebra \ (lambda_system \ M f)" apply (auto simp add: algebra_iff_Un) apply (metis lambda_system_sets set_mp sets_into_space) apply (metis lambda_system_empty) @@ -238,117 +163,62 @@ done lemma (in algebra) lambda_system_strong_additive: - assumes z: "z \ sets M" and disj: "x \ y = {}" - and xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" + assumes z: "z \ 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" + have "(z \ (x \ y)) \ 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" +lemma (in algebra) lambda_system_additive: "additive (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)+ + and xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" + hence "x \ M" "y \ 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 ring_of_sets) disjointed_additive: - assumes f: "positive M f" "additive M f" and A: "range A \ sets M" "incseq A" - shows "(\i\n. f (disjointed A i)) = f (A n)" -proof (induct n) - case (Suc n) - then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" - by simp - also have "\ = f (A n \ disjointed A (Suc n))" - using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq) - also have "A n \ disjointed A (Suc n) = A (Suc n)" - using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq) - finally show ?case . -qed simp - lemma (in ring_of_sets) countably_subadditive_subadditive: assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y - assume x: "x \ sets M" and y: "y \ sets M" and "x \ y = {}" + assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) - hence "range (binaryset x y) \ sets M \ - (\i. binaryset x y i) \ sets M \ + hence "range (binaryset x y) \ M \ + (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" using cs by (auto simp add: countably_subadditive_def) - hence "{x,y,{}} \ sets M \ x \ y \ sets M \ + hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) \ (\ n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) \ f x + f y" using f x y by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed -lemma (in ring_of_sets) additive_sum: - fixes A:: "nat \ 'a set" - assumes f: "positive M f" and ad: "additive M f" and "finite S" - and A: "range A \ sets M" - and disj: "disjoint_family_on A S" - shows "(\i\S. f (A i)) = f (\i\S. A i)" -using `finite S` disj proof induct - case empty show ?case using f by (simp add: positive_def) -next - case (insert s S) - then have "A s \ (\i\S. A i) = {}" - by (auto simp add: disjoint_family_on_def neq_iff) - moreover - have "A s \ sets M" using A by blast - moreover have "(\i\S. A i) \ sets M" - using A `finite S` by auto - moreover - ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" - using ad UNION_in_sets A by (auto simp add: additive_def) - with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] - by (auto simp add: additive_def subset_insertI) -qed - -lemma (in algebra) increasing_additive_bound: - fixes A:: "nat \ 'a set" and f :: "'a set \ ereal" - 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" - shows "(\i. f (A i)) \ f (space M)" -proof (safe intro!: suminf_bound) - fix N - note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. f (space M)" using space_closed A - by (intro increasingD[OF inc] finite_UN) auto - finally show "(\i f (space M)" by simp -qed (insert f A, auto simp: positive_def) - lemma lambda_system_increasing: - "increasing M f \ increasing (M (|sets := lambda_system M f|)) f" + "increasing M f \ increasing (lambda_system \ M f) f" by (simp add: increasing_def lambda_system_def) lemma lambda_system_positive: - "positive M f \ positive (M (|sets := lambda_system M f|)) f" + "positive M f \ positive (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 \ ereal" - assumes f: "positive M f" and a: "a \ sets M" - and A: "range A \ lambda_system M f" + assumes f: "positive M f" and a: "a \ M" + and A: "range A \ lambda_system \ M f" and disj: "disjoint_family A" shows "(\i = 0..A i)) = f (a \ (\i\{0.. UNION {0.. lambda_system M f" using A + have 3: "A n \ lambda_system \ M f" using A by blast - interpret l: algebra "M\sets := lambda_system M f\" + interpret l: algebra \ "lambda_system \ M f" using f by (rule lambda_system_algebra) - have 4: "UNION {0.. lambda_system M f" + have 4: "UNION {0.. lambda_system \ M f" using A l.UNION_in_sets by simp from Suc.hyps show ?case by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4]) @@ -369,23 +239,23 @@ lemma (in sigma_algebra) lambda_system_caratheodory: assumes oms: "outer_measure_space M f" - and A: "range A \ lambda_system M f" + and A: "range A \ lambda_system \ M f" and disj: "disjoint_family A" - shows "(\i. A i) \ lambda_system M f \ (\i. f (A i)) = f (\i. A i)" + shows "(\i. A i) \ lambda_system \ M f \ (\i. f (A i)) = f (\i. A i)" proof - 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" by (metis countably_subadditive_subadditive csa pos) - have A': "range A \ sets (M(|sets := lambda_system M f|))" using A - by simp - interpret ls: algebra "M\sets := lambda_system M f\" + have A': "\S. A`S \ (lambda_system \ M f)" using A + by auto + interpret ls: algebra \ "lambda_system \ M f" using pos by (rule lambda_system_algebra) - have A'': "range A \ sets M" + have A'': "range A \ M" by (metis A image_subset_iff lambda_system_sets) - have U_in: "(\i. A i) \ sets M" + have U_in: "(\i. A i) \ M" by (metis A'' countable_UN) have U_eq: "f (\i. A i) = (\i. f (A i))" proof (rule antisym) @@ -396,22 +266,22 @@ show "(\i. f (A i)) \ f (\i. A i)" using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] using A'' - by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN) + by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] countable_UN) qed { fix a - assume a [iff]: "a \ sets M" + assume a [iff]: "a \ M" have "f (a \ (\i. A i)) + f (a - (\i. A i)) = f a" proof - show ?thesis proof (rule antisym) - have "range (\i. a \ A i) \ sets M" using A'' + have "range (\i. a \ A i) \ M" using A'' by blast moreover have "disjoint_family (\i. a \ A i)" using disj by (auto simp add: disjoint_family_on_def) moreover - have "a \ (\i. A i) \ sets M" + have "a \ (\i. A i) \ M" by (metis Int U_in a) ultimately have "f (a \ (\i. A i)) \ (\i. f (a \ A i))" @@ -424,11 +294,11 @@ have "(\i. f (a \ A i)) + f (a - (\i. A i)) \ f a" proof (intro suminf_bound_add allI) fix n - have UNION_in: "(\i\{0.. sets M" + have UNION_in: "(\i\{0.. M" by (metis A'' UNION_in_sets) have le_fa: "f (UNION {0.. a) \ f a" using A'' by (blast intro: increasingD [OF inc] A'' UNION_in_sets) - have ls: "(\i\{0.. lambda_system M f" + have ls: "(\i\{0.. lambda_system \ M f" using ls.UNION_in_sets by (simp add: A) hence eq_fa: "f a = f (a \ (\i\{0..i\{0..i A i)) + f (a - (\i. A i)) \ f a" by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) next - have "\i. a \ A i \ sets M" using A'' by auto + have "\i. a \ A i \ M" using A'' by auto then show "\i. 0 \ f (a \ A i)" using pos[unfolded positive_def] by auto - have "\i. a - (\i. A i) \ sets M" using A'' by auto + have "\i. a - (\i. A i) \ M" using A'' by auto then have "\i. 0 \ f (a - (\i. A i))" using pos[unfolded positive_def] by auto then show "f (a - (\i. A i)) \ -\" by auto qed @@ -460,66 +330,29 @@ lemma (in sigma_algebra) caratheodory_lemma: assumes oms: "outer_measure_space M f" - shows "measure_space \ space = space M, sets = lambda_system M f, measure = f \" - (is "measure_space ?M") + defines "L \ lambda_system \ M f" + shows "measure_space \ L f" proof - have pos: "positive M f" by (metis oms outer_measure_space_def) - have alg: "algebra ?M" + have alg: "algebra \ L" using lambda_system_algebra [of f, OF pos] - by (simp add: algebra_iff_Un) + by (simp add: algebra_iff_Un L_def) then - have "sigma_algebra ?M" + have "sigma_algebra \ L" using lambda_system_caratheodory [OF oms] - by (simp add: sigma_algebra_disjoint_iff) + by (simp add: sigma_algebra_disjoint_iff L_def) moreover - have "measure_space_axioms ?M" + have "countably_additive L f" "positive L f" using pos lambda_system_caratheodory [OF oms] - by (simp add: measure_space_axioms_def positive_def lambda_system_sets - countably_additive_def o_def) + by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def) ultimately show ?thesis - by (simp add: measure_space_def) -qed - -lemma (in ring_of_sets) additive_increasing: - assumes posf: "positive M f" and addf: "additive M f" - shows "increasing M f" -proof (auto simp add: increasing_def) - fix x y - assume xy: "x \ sets M" "y \ sets M" "x \ y" - then have "y - x \ sets M" by auto - then have "0 \ f (y-x)" using posf[unfolded positive_def] by auto - then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono) auto - also have "... = f (x \ (y-x))" using addf - by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) - also have "... = f y" - by (metis Un_Diff_cancel Un_absorb1 xy(3)) - finally show "f x \ f y" by simp -qed - -lemma (in ring_of_sets) countably_additive_additive: - assumes posf: "positive M f" and ca: "countably_additive M f" - shows "additive M f" -proof (auto simp add: additive_def) - fix x y - assume x: "x \ sets M" and y: "y \ sets M" and "x \ y = {}" - hence "disjoint_family (binaryset x y)" - by (auto simp add: disjoint_family_on_def binaryset_def) - hence "range (binaryset x y) \ sets M \ - (\i. binaryset x y i) \ sets M \ - f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" - using ca - by (simp add: countably_additive_def) - hence "{x,y,{}} \ sets M \ x \ y \ sets M \ - f (x \ y) = (\n. f (binaryset x y n))" - by (simp add: range_binaryset_eq UN_binaryset_eq) - thus "f (x \ y) = f x + f y" using posf x y - by (auto simp add: Un suminf_binaryset_eq positive_def) + using pos by (simp add: measure_space_def) qed lemma inf_measure_nonempty: - assumes f: "positive M f" and b: "b \ sets M" and a: "a \ b" "{} \ sets M" + assumes f: "positive M f" and b: "b \ M" and a: "a \ b" "{} \ M" shows "f b \ measure_set M f a" proof - let ?A = "\i::nat. (if i = 0 then b else {})" @@ -534,14 +367,14 @@ lemma (in ring_of_sets) inf_measure_agrees: assumes posf: "positive M f" and ca: "countably_additive M f" - and s: "s \ sets M" + and s: "s \ M" shows "Inf (measure_set M f s) = f s" unfolding Inf_ereal_def proof (safe intro!: Greatest_equality) fix z assume z: "z \ measure_set M f s" from this obtain A where - A: "range A \ sets M" and disj: "disjoint_family A" + A: "range A \ M" and disj: "disjoint_family A" and "s \ (\x. A x)" and si: "(\i. f (A i)) = z" by (auto simp add: measure_set_def comp_def) hence seq: "s = (\i. A i \ s)" by blast @@ -549,11 +382,11 @@ by (metis additive_increasing ca countably_additive_additive posf) have sums: "(\i. f (A i \ s)) = f (\i. A i \ s)" proof (rule ca[unfolded countably_additive_def, rule_format]) - show "range (\n. A n \ s) \ sets M" using A s + show "range (\n. A n \ s) \ M" using A s by blast show "disjoint_family (\n. A n \ s)" using disj by (auto simp add: disjoint_family_on_def) - show "(\i. A i \ s) \ sets M" using A s + show "(\i. A i \ s) \ M" using A s by (metis UN_extend_simps(4) s seq) qed hence "f s = (\i. f (A i \ s))" @@ -562,7 +395,7 @@ proof (rule suminf_le_pos) fix n show "f (A n \ s) \ f (A n)" using A s by (force intro: increasingD [OF inc]) - fix N have "A N \ s \ sets M" using A s by auto + fix N have "A N \ s \ M" using A s by auto then show "0 \ f (A N \ s)" using posf unfolding positive_def by auto qed also have "... = z" by (rule si) @@ -578,7 +411,7 @@ assumes posf: "positive M f" "r \ measure_set M f X" shows "0 \ r" proof - - obtain A where "range A \ sets M" and r: "r = (\i. f (A i))" + obtain A where "range A \ M" and r: "r = (\i. f (A i))" using `r \ measure_set M f X` unfolding measure_set_def by auto then show "0 \ r" using posf unfolding r positive_def by (intro suminf_0_le) auto @@ -593,26 +426,25 @@ qed lemma inf_measure_empty: - assumes posf: "positive M f" and "{} \ sets M" + assumes posf: "positive M f" and "{} \ 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` + by (metis complete_lattice_class.Inf_lower `{} \ M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) qed (rule inf_measure_pos[OF posf]) lemma (in ring_of_sets) inf_measure_positive: - assumes p: "positive M f" and "{} \ sets M" + assumes p: "positive M f" and "{} \ M" shows "positive M (\x. Inf (measure_set M f x))" proof (unfold positive_def, intro conjI ballI) show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto - fix A assume "A \ sets M" + fix A assume "A \ M" qed (rule inf_measure_pos[OF p]) lemma (in ring_of_sets) inf_measure_increasing: assumes posf: "positive M f" - shows "increasing \ space = space M, sets = Pow (space M) \ - (\x. Inf (measure_set M f x))" + shows "increasing (Pow \) (\x. Inf (measure_set M f x))" apply (clarsimp simp add: increasing_def) apply (rule complete_lattice_class.Inf_greatest) apply (rule complete_lattice_class.Inf_lower) @@ -621,13 +453,13 @@ lemma (in ring_of_sets) inf_measure_le: assumes posf: "positive M f" and inc: "increasing M f" - and x: "x \ {r . \A. range A \ sets M \ s \ (\i. A i) \ (\i. f (A i)) = r}" + and x: "x \ {r . \A. range A \ M \ s \ (\i. A i) \ (\i. f (A i)) = r}" shows "Inf (measure_set M f s) \ x" proof - - obtain A where A: "range A \ sets M" and ss: "s \ (\i. A i)" + obtain A where A: "range A \ M" and ss: "s \ (\i. A i)" and xeq: "(\i. f (A i)) = x" using x by auto - have dA: "range (disjointed A) \ sets M" + have dA: "range (disjointed A) \ M" by (metis A range_disjointed_sets) have "\n. f (disjointed A n) \ f (A n)" by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def) @@ -648,8 +480,8 @@ lemma (in ring_of_sets) inf_measure_close: fixes e :: ereal - assumes posf: "positive M f" and e: "0 < e" and ss: "s \ (space M)" and "Inf (measure_set M f s) \ \" - shows "\A. range A \ sets M \ disjoint_family A \ s \ (\i. A i) \ + assumes posf: "positive M f" and e: "0 < e" and ss: "s \ (\)" and "Inf (measure_set M f s) \ \" + shows "\A. range A \ M \ disjoint_family A \ s \ (\i. A i) \ (\i. f (A i)) \ Inf (measure_set M f s) + e" proof - from `Inf (measure_set M f s) \ \` have fin: "\Inf (measure_set M f s)\ \ \" @@ -662,22 +494,21 @@ lemma (in ring_of_sets) inf_measure_countably_subadditive: 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))" + shows "countably_subadditive (Pow \) (\x. Inf (measure_set M f x))" proof (simp add: countably_subadditive_def, safe) fix A :: "nat \ 'a set" let ?outer = "\B. Inf (measure_set M f B)" - assume A: "range A \ Pow (space M)" + assume A: "range A \ Pow (\)" and disj: "disjoint_family A" - and sb: "(\i. A i) \ space M" + and sb: "(\i. A i) \ \" { fix e :: ereal assume e: "0 < e" and "\i. ?outer (A i) \ \" - hence "\BB. \n. range (BB n) \ sets M \ disjoint_family (BB n) \ + hence "\BB. \n. range (BB n) \ M \ disjoint_family (BB n) \ A n \ (\i. BB n i) \ (\i. f (BB n i)) \ ?outer (A n) + e * (1/2)^(Suc n)" apply (safe intro!: choice inf_measure_close [of f, OF posf]) using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def) then obtain BB - where BB: "\n. (range (BB n) \ sets M)" + where BB: "\n. (range (BB n) \ M)" and disjBB: "\n. disjoint_family (BB n)" and sbBB: "\n. A n \ (\i. BB n i)" and BBle: "\n. (\i. f (BB n i)) \ ?outer (A n) + e * (1/2)^(Suc n)" @@ -697,7 +528,7 @@ finally show ?thesis . qed def C \ "(split BB) o prod_decode" - have C: "!!n. C n \ sets M" + have C: "!!n. C n \ M" apply (rule_tac p="prod_decode n" in PairE) apply (simp add: C_def) apply (metis BB subsetD rangeI) @@ -744,9 +575,8 @@ qed lemma (in ring_of_sets) inf_measure_outer: - "\ positive M f ; increasing M f \ - \ outer_measure_space \ space = space M, sets = Pow (space M) \ - (\x. Inf (measure_set M f x))" + "\ positive M f ; increasing M f \ \ + outer_measure_space (Pow \) (\x. Inf (measure_set M f x))" using inf_measure_pos[of M f] by (simp add: outer_measure_space_def inf_measure_empty inf_measure_increasing inf_measure_countably_subadditive positive_def) @@ -754,14 +584,13 @@ lemma (in ring_of_sets) algebra_subset_lambda_system: assumes posf: "positive M f" and inc: "increasing M f" and add: "additive M f" - shows "sets M \ lambda_system \ space = space M, sets = Pow (space M) \ - (\x. Inf (measure_set M f x))" + shows "M \ lambda_system \ (Pow \) (\x. Inf (measure_set M f x))" proof (auto dest: sets_into_space simp add: algebra.lambda_system_eq [OF algebra_Pow]) fix x s - assume x: "x \ sets M" - and s: "s \ space M" - have [simp]: "!!x. x \ sets M \ s \ (space M - x) = s-x" using s + assume x: "x \ M" + and s: "s \ \" + have [simp]: "!!x. x \ M \ s \ (\ - x) = s-x" using s by blast have "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) \ Inf (measure_set M f s)" @@ -774,7 +603,7 @@ show ?thesis proof (rule complete_lattice_class.Inf_greatest) fix r assume "r \ measure_set M f s" - then obtain A where A: "disjoint_family A" "range A \ sets M" "s \ (\i. A i)" + then obtain A where A: "disjoint_family A" "range A \ M" "s \ (\i. A i)" and r: "r = (\i. f (A i))" unfolding measure_set_def by auto have "Inf (measure_set M f (s \ x)) \ (\i. f (A i \ x))" unfolding measure_set_def @@ -822,34 +651,31 @@ qed lemma measure_down: - "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) + "measure_space \ N \ \ sigma_algebra \ M \ M \ N \ measure_space \ M \" + by (simp add: measure_space_def positive_def countably_additive_def) blast theorem (in ring_of_sets) caratheodory: assumes posf: "positive M f" and ca: "countably_additive M f" - shows "\\ :: 'a set \ ereal. (\s \ sets M. \ s = f s) \ - measure_space \ space = space M, sets = sets (sigma M), measure = \ \" + shows "\\ :: 'a set \ ereal. (\s \ M. \ s = f s) \ measure_space \ (sigma_sets \ 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, measure = ?infm\" + def ls \ "lambda_system \ (Pow \) ?infm" + have mls: "measure_space \ 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, measure = ?infm|)" + hence sls: "sigma_algebra \ ls" by (simp add: measure_space_def) - have "sets M \ ls" + have "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"] + hence sgs_sb: "sigma_sets (\) (M) \ ls" + using sigma_algebra.sigma_sets_subset [OF sls, of "M"] by simp - have "measure_space \ space = space M, sets = sets (sigma M), measure = ?infm \" - unfolding sigma_def + have "measure_space \ (sigma_sets \ M) ?infm" 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] @@ -861,12 +687,12 @@ lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: assumes f: "positive M f" "additive M f" shows "countably_additive M f \ - (\A. range A \ sets M \ incseq A \ (\i. A i) \ sets M \ (\i. f (A i)) ----> f (\i. A i))" + (\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) ----> f (\i. A i))" unfolding countably_additive_def proof safe - assume count_sum: "\A. range A \ sets M \ disjoint_family A \ UNION UNIV A \ sets M \ (\i. f (A i)) = f (UNION UNIV A)" - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "incseq A" "(\i. A i) \ sets M" - then have dA: "range (disjointed A) \ sets M" by (auto simp: range_disjointed_sets) + assume count_sum: "\A. range A \ M \ disjoint_family A \ UNION UNIV A \ M \ (\i. f (A i)) = f (UNION UNIV A)" + fix A :: "nat \ 'a set" assume A: "range A \ M" "incseq A" "(\i. A i) \ M" + then have dA: "range (disjointed A) \ M" by (auto simp: range_disjointed_sets) with count_sum[THEN spec, of "disjointed A"] A(3) have f_UN: "(\i. f (disjointed A i)) = f (\i. A i)" by (auto simp: UN_disjointed_eq disjoint_family_disjointed) @@ -880,12 +706,12 @@ using disjointed_additive[OF f A(1,2)] . ultimately show "(\i. f (A i)) ----> f (\i. A i)" by simp next - assume cont: "\A. range A \ sets M \ incseq A \ (\i. A i) \ sets M \ (\i. f (A i)) ----> f (\i. A i)" - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" "(\i. A i) \ sets M" + assume cont: "\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) ----> f (\i. A i)" + fix A :: "nat \ 'a set" assume A: "range A \ M" "disjoint_family A" "(\i. A i) \ M" have *: "(\n. (\i\n. A i)) = (\i. A i)" by auto have "(\n. f (\i\n. A i)) ----> f (\i. A i)" proof (unfold *[symmetric], intro cont[rule_format]) - show "range (\i. \ i\i. A i) \ sets M" "(\i. \ i\i. A i) \ sets M" + show "range (\i. \ i\i. A i) \ M" "(\i. \ i\i. A i) \ M" using A * by auto qed (force intro!: incseq_SucI) moreover have "\n. f (\i\n. A i) = (\i\n. f (A i))" @@ -901,18 +727,18 @@ lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: assumes f: "positive M f" "additive M f" - shows "(\A. range A \ sets M \ decseq A \ (\i. A i) \ sets M \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> f (\i. A i)) - \ (\A. range A \ sets M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> 0)" + shows "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> f (\i. A i)) + \ (\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> 0)" proof safe - assume cont: "(\A. range A \ sets M \ decseq A \ (\i. A i) \ sets M \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> f (\i. A i))" - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" + assume cont: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> f (\i. A i))" + fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[THEN spec, of A] show "(\i. f (A i)) ----> 0" using `positive M f`[unfolded positive_def] by auto next - assume cont: "\A. range A \ sets M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> 0" - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "decseq A" "(\i. A i) \ sets M" "\i. f (A i) \ \" + assume cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) ----> 0" + fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" - have f_mono: "\a b. a \ sets M \ b \ sets M \ a \ b \ f a \ f b" + have f_mono: "\a b. a \ M \ b \ M \ a \ b \ f a \ f b" using additive_increasing[OF f] unfolding increasing_def by simp have decseq_fA: "decseq (\i. f (A i))" @@ -932,7 +758,7 @@ note f_fin = this have "(\i. f (A i - (\i. A i))) ----> 0" proof (intro cont[rule_format, OF _ decseq _ f_fin]) - show "range (\i. A i - (\i. A i)) \ sets M" "(\i. A i - (\i. A i)) = {}" + show "range (\i. A i - (\i. A i)) \ M" "(\i. A i - (\i. A i)) = {}" using A by auto qed from INF_Lim_ereal[OF decseq_f this] @@ -956,17 +782,17 @@ qed lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) -lemma positiveD2: "positive M f \ A \ sets M \ 0 \ f A" by (auto simp: positive_def) +lemma positiveD2: "positive M f \ A \ M \ 0 \ f A" by (auto simp: positive_def) lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: - assumes f: "positive M f" "additive M f" "\A\sets M. f A \ \" - assumes cont: "\A. range A \ sets M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" - assumes A: "range A \ sets M" "incseq A" "(\i. A i) \ sets M" + assumes f: "positive M f" "additive M f" "\A\M. f A \ \" + assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" + assumes A: "range A \ M" "incseq A" "(\i. A i) \ M" shows "(\i. f (A i)) ----> f (\i. A i)" proof - - have "\A\sets M. \x. f A = ereal x" + have "\A\M. \x. f A = ereal x" proof - fix A assume "A \ sets M" with f show "\x. f A = ereal x" + fix A assume "A \ M" with f show "\x. f A = ereal x" unfolding positive_def by (cases "f A") auto qed from bchoice[OF this] guess f' .. note f' = this[rule_format] @@ -991,20 +817,19 @@ qed lemma (in ring_of_sets) empty_continuous_imp_countably_additive: - assumes f: "positive M f" "additive M f" and fin: "\A\sets M. f A \ \" - assumes cont: "\A. range A \ sets M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" + assumes f: "positive M f" "additive M f" and fin: "\A\M. f A \ \" + assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" shows "countably_additive M f" using countably_additive_iff_continuous_from_below[OF f] using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast lemma (in ring_of_sets) caratheodory_empty_continuous: - assumes f: "positive M f" "additive M f" and fin: "\A. A \ sets M \ f A \ \" - assumes cont: "\A. range A \ sets M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" - shows "\\ :: 'a set \ ereal. (\s \ sets M. \ s = f s) \ - measure_space \ space = space M, sets = sets (sigma M), measure = \ \" + assumes f: "positive M f" "additive M f" and fin: "\A. A \ M \ f A \ \" + assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" + shows "\\ :: 'a set \ ereal. (\s \ M. \ s = f s) \ measure_space \ (sigma_sets \ M) \" proof (intro caratheodory empty_continuous_imp_countably_additive f) - show "\A\sets M. f A \ \" using fin by auto + show "\A\M. f A \ \" using fin by auto qed (rule cont) end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Complete_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -6,218 +6,231 @@ imports Lebesgue_Integration begin -locale completeable_measure_space = measure_space - -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 + "split_completion M A p = (if A \ sets M then p = (A, {}) else + \N'. A = fst p \ snd p \ fst p \ snd p = {} \ fst p \ sets M \ snd p \ N' \ N' \ null_sets M)" -definition (in completeable_measure_space) - "null_part A = snd (Eps (split_completion A))" - -abbreviation (in completeable_measure_space) "\' A \ \ (main_part A)" +definition + "main_part M A = fst (Eps (split_completion M 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' }, - measure = \', - \ = more M \" - +definition + "null_part M A = snd (Eps (split_completion M A))" -lemma (in completeable_measure_space) space_completion[simp]: - "space completion = space M" unfolding completion_def by simp - -lemma (in completeable_measure_space) sets_completionE: - assumes "A \ sets completion" - obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets" "S \ sets M" - using assms unfolding completion_def by auto +definition completion :: "'a measure \ 'a measure" where + "completion M = measure_of (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } + (emeasure M \ main_part M)" -lemma (in completeable_measure_space) sets_completionI: - assumes "A = S \ N" "N \ N'" "N' \ null_sets" "S \ sets M" - shows "A \ sets completion" - using assms unfolding completion_def by auto +lemma completion_into_space: + "{ S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } \ Pow (space M)" + using sets_into_space by auto -lemma (in completeable_measure_space) sets_completionI_sets[intro]: - "A \ sets M \ A \ sets completion" - unfolding completion_def by force +lemma space_completion[simp]: "space (completion M) = space M" + unfolding completion_def using space_measure_of[OF completion_into_space] by simp -lemma (in completeable_measure_space) null_sets_completion: - assumes "N' \ null_sets" "N \ N'" shows "N \ sets completion" - apply(rule sets_completionI[of N "{}" N N']) +lemma completionI: + assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + shows "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" using assms by auto -sublocale completeable_measure_space \ completion!: sigma_algebra completion -proof (unfold sigma_algebra_iff2, safe) - fix A x assume "A \ sets completion" "x \ A" - with sets_into_space show "x \ space completion" - by (auto elim!: sets_completionE) +lemma completionE: + assumes "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" + obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + using assms by auto + +lemma sigma_algebra_completion: + "sigma_algebra (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" + (is "sigma_algebra _ ?A") + unfolding sigma_algebra_iff2 +proof (intro conjI ballI allI impI) + show "?A \ Pow (space M)" + using sets_into_space by auto next - fix A assume "A \ sets completion" - from this[THEN sets_completionE] guess S N N' . note A = this - let ?C = "space completion" - show "?C - A \ sets completion" using A - by (intro sets_completionI[of _ "(?C - S) \ (?C - N')" "(?C - S) \ N' \ (?C - N)"]) - auto + show "{} \ ?A" by auto next - fix A ::"nat \ 'a set" assume A: "range A \ sets completion" - then have "\n. \S N N'. A n = S \ N \ S \ sets M \ N' \ null_sets \ N \ N'" - unfolding completion_def by (auto simp: image_subset_iff) + let ?C = "space M" + fix A assume "A \ ?A" from completionE[OF this] guess S N N' . + then show "space M - A \ ?A" + by (intro completionI[of _ "(?C - S) \ (?C - N')" "(?C - S) \ N' \ (?C - N)"]) auto +next + fix A :: "nat \ 'a set" assume A: "range A \ ?A" + then have "\n. \S N N'. A n = S \ N \ S \ sets M \ N' \ null_sets M \ N \ N'" + by (auto simp: image_subset_iff) from choice[OF this] guess S .. from choice[OF this] guess N .. from choice[OF this] guess N' .. - then show "UNION UNIV A \ sets completion" + then show "UNION UNIV A \ ?A" using null_sets_UN[of N'] - by (intro sets_completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) - auto -qed auto + by (intro completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) auto +qed + +lemma sets_completion: + "sets (completion M) = { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" + using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def) + +lemma sets_completionE: + assumes "A \ sets (completion M)" + obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + using assms unfolding sets_completion by auto + +lemma sets_completionI: + assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + shows "A \ sets (completion M)" + using assms unfolding sets_completion by auto + +lemma sets_completionI_sets[intro, simp]: + "A \ sets M \ A \ sets (completion M)" + unfolding sets_completion by force -lemma (in completeable_measure_space) split_completion: - assumes "A \ sets completion" - shows "split_completion A (main_part A, null_part A)" - unfolding main_part_def null_part_def -proof (rule someI2_ex) - from assms[THEN sets_completionE] guess S N N' . note A = this - let ?P = "(S, N - S)" - show "\p. split_completion A p" - unfolding split_completion_def using A - proof (intro exI conjI) - show "A = fst ?P \ snd ?P" using A by auto - show "snd ?P \ N'" using A by auto +lemma null_sets_completion: + assumes "N' \ null_sets M" "N \ N'" shows "N \ sets (completion M)" + using assms by (intro sets_completionI[of N "{}" N N']) auto + +lemma split_completion: + assumes "A \ sets (completion M)" + shows "split_completion M A (main_part M A, null_part M A)" +proof cases + assume "A \ sets M" then show ?thesis + by (simp add: split_completion_def[abs_def] main_part_def null_part_def) +next + assume nA: "A \ sets M" + show ?thesis + unfolding main_part_def null_part_def if_not_P[OF nA] + proof (rule someI2_ex) + from assms[THEN sets_completionE] guess S N N' . note A = this + let ?P = "(S, N - S)" + show "\p. split_completion M A p" + unfolding split_completion_def if_not_P[OF nA] using A + proof (intro exI conjI) + show "A = fst ?P \ snd ?P" using A by auto + show "snd ?P \ N'" using A by auto + qed auto qed auto -qed auto +qed -lemma (in completeable_measure_space) - assumes "S \ sets completion" - shows main_part_sets[intro, simp]: "main_part S \ sets M" - and main_part_null_part_Un[simp]: "main_part S \ null_part S = S" - and main_part_null_part_Int[simp]: "main_part S \ null_part S = {}" - using split_completion[OF assms] by (auto simp: split_completion_def) +lemma + assumes "S \ sets (completion M)" + shows main_part_sets[intro, simp]: "main_part M S \ sets M" + and main_part_null_part_Un[simp]: "main_part M S \ null_part M S = S" + and main_part_null_part_Int[simp]: "main_part M S \ null_part M S = {}" + using split_completion[OF assms] + by (auto simp: split_completion_def split: split_if_asm) -lemma (in completeable_measure_space) null_part: - assumes "S \ sets completion" shows "\N. N\null_sets \ null_part S \ N" - using split_completion[OF assms] by (auto simp: split_completion_def) +lemma main_part[simp]: "S \ sets M \ main_part M S = S" + using split_completion[of S M] + by (auto simp: split_completion_def split: split_if_asm) -lemma (in completeable_measure_space) null_part_sets[intro, simp]: - assumes "S \ sets M" shows "null_part S \ sets M" "\ (null_part S) = 0" +lemma null_part: + assumes "S \ sets (completion M)" shows "\N. N\null_sets M \ null_part M S \ N" + using split_completion[OF assms] by (auto simp: split_completion_def split: split_if_asm) + +lemma null_part_sets[intro, simp]: + assumes "S \ sets M" shows "null_part M S \ sets M" "emeasure M (null_part M S) = 0" proof - - have S: "S \ sets completion" using assms by auto - have "S - main_part S \ sets M" using assms by auto + have S: "S \ sets (completion M)" using assms by auto + have "S - main_part M S \ sets M" using assms by auto moreover from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S] - have "S - main_part S = null_part S" by auto - ultimately show sets: "null_part S \ sets M" by auto + have "S - main_part M S = null_part M S" by auto + ultimately show sets: "null_part M S \ sets M" by auto from null_part[OF S] guess N .. - with measure_eq_0[of N "null_part S"] sets - show "\ (null_part S) = 0" by auto -qed - -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 "\ = \' S" - using S assms measure_additive[of "main_part S" "null_part S"] - by (auto simp: measure_additive) - finally show ?thesis by simp + with emeasure_eq_0[of N _ "null_part M S"] sets + show "emeasure M (null_part M S) = 0" by auto qed -lemma (in completeable_measure_space) sets_completionI_sub: - assumes N: "N' \ null_sets" "N \ N'" - shows "N \ sets completion" - using assms by (intro sets_completionI[of _ "{}" N N']) auto - -lemma (in completeable_measure_space) \_main_part_UN: +lemma emeasure_main_part_UN: fixes S :: "nat \ 'a set" - assumes "range S \ sets completion" - shows "\' (\i. (S i)) = \ (\i. main_part (S i))" + assumes "range S \ sets (completion M)" + shows "emeasure M (main_part M (\i. (S i))) = emeasure M (\i. main_part M (S i))" proof - - have S: "\i. S i \ sets completion" using assms by auto - then have UN: "(\i. S i) \ sets completion" by auto - have "\i. \N. N \ null_sets \ null_part (S i) \ N" + have S: "\i. S i \ sets (completion M)" using assms by auto + then have UN: "(\i. S i) \ sets (completion M)" by auto + have "\i. \N. N \ null_sets M \ null_part M (S i) \ N" using null_part[OF S] by auto from choice[OF this] guess N .. note N = this - then have UN_N: "(\i. N i) \ null_sets" by (intro null_sets_UN) auto - have "(\i. S i) \ sets completion" using S by auto + then have UN_N: "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto + have "(\i. S i) \ sets (completion M)" using S by auto from null_part[OF this] guess N' .. note N' = this let ?N = "(\i. N i) \ N'" - have null_set: "?N \ null_sets" using N' UN_N by (intro nullsets.Un) auto - have "main_part (\i. S i) \ ?N = (main_part (\i. S i) \ null_part (\i. S i)) \ ?N" + have null_set: "?N \ null_sets M" using N' UN_N by (intro null_sets.Un) auto + have "main_part M (\i. S i) \ ?N = (main_part M (\i. S i) \ null_part M (\i. S i)) \ ?N" using N' by auto - also have "\ = (\i. main_part (S i) \ null_part (S i)) \ ?N" + also have "\ = (\i. main_part M (S i) \ null_part M (S i)) \ ?N" unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto - also have "\ = (\i. main_part (S i)) \ ?N" + also have "\ = (\i. main_part M (S i)) \ ?N" using N by auto - finally have *: "main_part (\i. S i) \ ?N = (\i. main_part (S i)) \ ?N" . - have "\ (main_part (\i. S i)) = \ (main_part (\i. S i) \ ?N)" - using null_set UN by (intro measure_Un_null_set[symmetric]) auto - also have "\ = \ ((\i. main_part (S i)) \ ?N)" + finally have *: "main_part M (\i. S i) \ ?N = (\i. main_part M (S i)) \ ?N" . + have "emeasure M (main_part M (\i. S i)) = emeasure M (main_part M (\i. S i) \ ?N)" + using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto + also have "\ = emeasure M ((\i. main_part M (S i)) \ ?N)" unfolding * .. - also have "\ = \ (\i. main_part (S i))" - using null_set S by (intro measure_Un_null_set) auto + also have "\ = emeasure M (\i. main_part M (S i))" + using null_set S by (intro emeasure_Un_null_set) auto finally show ?thesis . qed -lemma (in completeable_measure_space) \_main_part_Un: - assumes S: "S \ sets completion" and T: "T \ sets completion" - shows "\' (S \ T) = \ (main_part S \ main_part T)" -proof - - have UN: "(\i. binary (main_part S) (main_part T) i) = (\i. main_part (binary S T i))" - unfolding binary_def by (auto split: split_if_asm) - show ?thesis - using \_main_part_UN[of "binary S T"] assms - unfolding range_binary_eq Un_range_binary UN by auto +lemma emeasure_completion[simp]: + assumes S: "S \ sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)" +proof (subst emeasure_measure_of[OF completion_def completion_into_space]) + let ?\ = "emeasure M \ main_part M" + show "S \ sets (completion M)" "?\ S = emeasure M (main_part M S) " using S by simp_all + show "positive (sets (completion M)) ?\" + by (simp add: positive_def emeasure_nonneg) + show "countably_additive (sets (completion M)) ?\" + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume A: "range A \ sets (completion M)" "disjoint_family A" + have "disjoint_family (\i. main_part M (A i))" + proof (intro disjoint_family_on_bisimulation[OF A(2)]) + fix n m assume "A n \ A m = {}" + then have "(main_part M (A n) \ null_part M (A n)) \ (main_part M (A m) \ null_part M (A m)) = {}" + using A by (subst (1 2) main_part_null_part_Un) auto + then show "main_part M (A n) \ main_part M (A m) = {}" by auto + qed + then have "(\n. emeasure M (main_part M (A n))) = emeasure M (\i. main_part M (A i))" + using A by (auto intro!: suminf_emeasure) + then show "(\n. ?\ (A n)) = ?\ (UNION UNIV A)" + by (simp add: completion_def emeasure_main_part_UN[OF A(1)]) + qed qed -sublocale completeable_measure_space \ completion!: measure_space completion - where "measure completion = \'" -proof - - show "measure_space completion" - proof - show "positive completion (measure completion)" - by (auto simp: completion_def positive_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 "(\n. measure completion (A n)) = \ (\i. main_part (A i))" - unfolding completion_def using A by (auto intro!: measure_countably_additive) - then show "(\n. measure completion (A n)) = measure completion (UNION UNIV A)" - by (simp add: completion_def \_main_part_UN[OF A(1)]) - qed - qed - show "measure completion = \'" unfolding completion_def by simp -qed +lemma emeasure_completion_UN: + "range S \ sets (completion M) \ + emeasure (completion M) (\i::nat. (S i)) = emeasure M (\i. main_part M (S i))" + by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN) -lemma (in completeable_measure_space) completion_ex_simple_function: - assumes f: "simple_function completion f" - shows "\f'. simple_function M f' \ (AE x. f x = f' x)" +lemma emeasure_completion_Un: + assumes S: "S \ sets (completion M)" and T: "T \ sets (completion M)" + shows "emeasure (completion M) (S \ T) = emeasure M (main_part M S \ main_part M T)" +proof (subst emeasure_completion) + have UN: "(\i. binary (main_part M S) (main_part M T) i) = (\i. main_part M (binary S T i))" + unfolding binary_def by (auto split: split_if_asm) + show "emeasure M (main_part M (S \ T)) = emeasure M (main_part M S \ main_part M T)" + using emeasure_main_part_UN[of "binary S T" M] assms + unfolding range_binary_eq Un_range_binary UN by auto +qed (auto intro: S T) + +lemma sets_completionI_sub: + assumes N: "N' \ null_sets M" "N \ N'" + shows "N \ sets (completion M)" + using assms by (intro sets_completionI[of _ "{}" N N']) auto + +lemma completion_ex_simple_function: + assumes f: "simple_function (completion M) f" + shows "\f'. simple_function M f' \ (AE x in M. 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)" - using completion.simple_functionD[OF f] - completion.simple_functionD[OF f] by simp_all - have "\x. \N. N \ null_sets \ null_part (?F x) \ N" + have F: "\x. ?F x \ sets (completion M)" and fin: "finite (f`space M)" + using simple_functionD[OF f] simple_functionD[OF f] by simp_all + have "\x. \N. N \ null_sets M \ null_part M (?F x) \ N" using F null_part by auto from choice[OF this] obtain N where - N: "\x. null_part (?F x) \ N x" "\x. N x \ null_sets" by auto + N: "\x. null_part M (?F x) \ N x" "\x. N x \ null_sets M" by auto let ?N = "\x\f`space M. N x" let ?f' = "\x. if x \ ?N then undefined else f x" - have sets: "?N \ null_sets" using N fin by (intro nullsets.finite_UN) auto + have sets: "?N \ null_sets M" using N fin by (intro null_sets.finite_UN) auto show ?thesis unfolding simple_function_def proof (safe intro!: exI[of _ ?f']) have "?f' ` space M \ f`space M \ {undefined}" by auto - from finite_subset[OF this] completion.simple_functionD(1)[OF f] + from finite_subset[OF this] simple_functionD(1)[OF f] show "finite (?f' ` space M)" by auto next fix x assume "x \ space M" @@ -225,13 +238,13 @@ (if x \ ?N then ?F undefined \ ?N else if f x = undefined then ?F (f x) \ ?N else ?F (f x) - ?N)" - using N(2) sets_into_space by (auto split: split_if_asm) + using N(2) sets_into_space by (auto split: split_if_asm simp: null_sets_def) moreover { fix y have "?F y \ ?N \ sets M" proof cases assume y: "y \ f`space M" - have "?F y \ ?N = (main_part (?F y) \ null_part (?F y)) \ ?N" + have "?F y \ ?N = (main_part M (?F y) \ null_part M (?F y)) \ ?N" using main_part_null_part_Un[OF F] by auto - also have "\ = main_part (?F y) \ ?N" + also have "\ = main_part M (?F y) \ ?N" using y N by auto finally show ?thesis using F sets by auto @@ -240,34 +253,34 @@ then show ?thesis using sets by auto qed } moreover { - have "?F (f x) - ?N = main_part (?F (f x)) \ null_part (?F (f x)) - ?N" + have "?F (f x) - ?N = main_part M (?F (f x)) \ null_part M (?F (f x)) - ?N" using main_part_null_part_Un[OF F] by auto - also have "\ = main_part (?F (f x)) - ?N" + also have "\ = main_part M (?F (f x)) - ?N" using N `x \ space M` by auto finally have "?F (f x) - ?N \ sets M" using F sets by auto } ultimately show "?f' -` {?f' x} \ space M \ sets M" by auto next - show "AE x. f x = ?f' x" + show "AE x in M. f x = ?f' x" by (rule AE_I', rule sets) auto qed qed -lemma (in completeable_measure_space) completion_ex_borel_measurable_pos: +lemma completion_ex_borel_measurable_pos: fixes g :: "'a \ ereal" - assumes g: "g \ borel_measurable completion" and "\x. 0 \ g x" - shows "\g'\borel_measurable M. (AE x. g x = g' x)" + assumes g: "g \ borel_measurable (completion M)" and "\x. 0 \ g x" + shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" proof - - from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this + from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this from this(1)[THEN completion_ex_simple_function] - have "\i. \f'. simple_function M f' \ (AE x. f i x = f' x)" .. + have "\i. \f'. simple_function M f' \ (AE x in M. f i x = f' x)" .. from this[THEN choice] obtain f' where sf: "\i. simple_function M (f' i)" and - AE: "\i. AE x. f i x = f' i x" by auto + AE: "\i. AE x in M. f i x = f' i x" by auto show ?thesis proof (intro bexI) from AE[unfolded AE_all_countable[symmetric]] - show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x") + show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x") proof (elim AE_mp, safe intro!: AE_I2) fix x assume eq: "\i. f i x = f' i x" moreover have "g x = (SUP i. f i x)" @@ -279,20 +292,20 @@ qed qed -lemma (in completeable_measure_space) completion_ex_borel_measurable: +lemma completion_ex_borel_measurable: fixes g :: "'a \ ereal" - assumes g: "g \ borel_measurable completion" - shows "\g'\borel_measurable M. (AE x. g x = g' x)" + assumes g: "g \ borel_measurable (completion M)" + shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" proof - - have "(\x. max 0 (g x)) \ borel_measurable completion" "\x. 0 \ max 0 (g x)" using g by auto + have "(\x. max 0 (g x)) \ borel_measurable (completion M)" "\x. 0 \ max 0 (g x)" using g by auto from completion_ex_borel_measurable_pos[OF this] guess g_pos .. moreover - have "(\x. max 0 (- g x)) \ borel_measurable completion" "\x. 0 \ max 0 (- g x)" using g by auto + have "(\x. max 0 (- g x)) \ borel_measurable (completion M)" "\x. 0 \ max 0 (- g x)" using g by auto from completion_ex_borel_measurable_pos[OF this] guess g_neg .. ultimately show ?thesis proof (safe intro!: bexI[of _ "\x. g_pos x - g_neg x"]) - show "AE x. max 0 (- g x) = g_neg x \ max 0 (g x) = g_pos x \ g x = g_pos x - g_neg x" + show "AE x in M. max 0 (- g x) = g_neg x \ max 0 (g x) = g_pos x \ g x = g_pos x - g_neg x" proof (intro AE_I2 impI) fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x" show "g x = g_pos x - g_neg x" unfolding g[symmetric] diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Conditional_Probability.thy --- a/src/HOL/Probability/Conditional_Probability.thy Mon Apr 23 12:23:23 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,161 +0,0 @@ -(* Title: HOL/Probability/Conditional_Probability.thy - Author: Johannes Hölzl, TU München -*) - -header {*Conditional probability*} - -theory Conditional_Probability -imports Probability_Measure Radon_Nikodym -begin - -section "Conditional Expectation and Probability" - -definition (in prob_space) - "conditional_expectation N X = (SOME Y. Y\borel_measurable N \ (\x. 0 \ Y x) - \ (\C\sets N. (\\<^isup>+x. Y x * indicator C x\M) = (\\<^isup>+x. X x * indicator C x\M)))" - -lemma (in prob_space) conditional_expectation_exists: - fixes X :: "'a \ ereal" and N :: "('a, 'b) measure_space_scheme" - assumes borel: "X \ borel_measurable M" "AE x. 0 \ X x" - and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" "\A. measure N A = \ A" - shows "\Y\borel_measurable N. (\x. 0 \ Y x) \ (\C\sets N. - (\\<^isup>+x. Y x * indicator C x \M) = (\\<^isup>+x. X x * indicator C x \M))" -proof - - 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. integral\<^isup>P M (?f A)" - - from measure_space_density[OF borel] - 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" "P.\ A = 0" - then have f_borel: "?f A \ borel_measurable M" "AE x. x \ A" - using borel N by (auto intro!: borel_measurable_indicator AE_not_in) - then show "?Q A = 0" - by (auto simp add: positive_integral_0_iff_AE) - qed - from P.Radon_Nikodym[OF Q this] - obtain Y where Y: "Y \ borel_measurable N" "\x. 0 \ Y 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,4,1)]) -qed - -lemma (in prob_space) - fixes X :: "'a \ ereal" and N :: "('a, 'b) measure_space_scheme" - assumes borel: "X \ borel_measurable M" "AE x. 0 \ X x" - 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 \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] - then show "conditional_expectation N X \ borel_measurable N" - unfolding conditional_expectation_def by (rule someI2_ex) blast - - from CE show "\C. C \ sets N \ ?eq C" - unfolding conditional_expectation_def by (rule someI2_ex) blast -qed - -lemma (in sigma_algebra) factorize_measurable_function_pos: - fixes Z :: "'a \ ereal" and Y :: "'a \ 'c" - assumes "sigma_algebra M'" and "Y \ measurable M M'" "Z \ borel_measurable M" - assumes Z: "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)" - shows "\g\borel_measurable M'. \x\space M. max 0 (Z x) = g (Y x)" -proof - - interpret M': sigma_algebra M' by fact - have Y: "Y \ space M \ space M'" using assms unfolding measurable_def by auto - from M'.sigma_algebra_vimage[OF this] - interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . - - from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this - - have "\i. \g. simple_function M' g \ (\x\space M. f i x = g (Y x))" - proof - fix i - from f(1)[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 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. simple_function M' g \ (\x\space M. f i x = g (Y x))" - proof (intro exI[of _ ?g] conjI ballI) - 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::ereal)" - unfolding indicator_def using B by auto - then show "f i x = ?g (Y x)" using `x \ space M` f(1)[of i] - by (subst va.simple_function_indicator_representation) auto - qed - qed - from choice[OF this] guess g .. note g = this - - show ?thesis - proof (intro ballI bexI) - show "(\x. SUP i. g i x) \ borel_measurable M'" - using g by (auto intro: M'.borel_measurable_simple_function) - fix x assume "x \ space M" - have "max 0 (Z x) = (SUP i. f i x)" using f by simp - also have "\ = (SUP i. g i (Y x))" - using g `x \ space M` by simp - finally show "max 0 (Z x) = (SUP i. g i (Y x))" . - qed -qed - -lemma (in sigma_algebra) factorize_measurable_function: - fixes Z :: "'a \ ereal" and Y :: "'a \ 'c" - assumes "sigma_algebra M'" and "Y \ measurable M M'" "Z \ borel_measurable M" - shows "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) - \ (\g\borel_measurable M'. \x\space M. Z x = g (Y x))" -proof safe - interpret M': sigma_algebra M' by fact - have Y: "Y \ space M \ space M'" using assms unfolding measurable_def by auto - from M'.sigma_algebra_vimage[OF this] - interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . - - { fix g :: "'c \ ereal" assume "g \ borel_measurable M'" - with M'.measurable_vimage_algebra[OF Y] - have "g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" - by (rule measurable_comp) - moreover assume "\x\space M. Z x = g (Y x)" - then have "Z \ borel_measurable (M'.vimage_algebra (space M) Y) \ - g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" - by (auto intro!: measurable_cong) - ultimately show "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" - by simp } - - assume Z: "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" - with assms have "(\x. - Z x) \ borel_measurable M" - "(\x. - Z x) \ borel_measurable (M'.vimage_algebra (space M) Y)" - by auto - from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this - from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this - let ?g = "\x. p x - n x" - show "\g\borel_measurable M'. \x\space M. Z x = g (Y x)" - proof (intro bexI ballI) - show "?g \ borel_measurable M'" using p n by auto - fix x assume "x \ space M" - then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)" - using p n by auto - then show "Z x = ?g (Y x)" - by (auto split: split_max) - qed -qed - -end \ No newline at end of file diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -8,6 +8,9 @@ imports Binary_Product_Measure begin +lemma split_const: "(\(i, j). c) = (\_. c)" + by auto + lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" unfolding Pi_def by auto @@ -34,9 +37,6 @@ notation (xsymbols) funcset_extensional (infixr "\\<^isub>E" 60) -lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" - by safe (auto simp add: extensional_def fun_eq_iff) - lemma extensional_insert[intro, simp]: assumes "a \ extensional (insert i I)" shows "a(i := b) \ extensional (insert i I)" @@ -86,7 +86,7 @@ "I \ J = {} \ restrict (merge I x J y) J = restrict y J" "J \ I = {} \ restrict (merge I x J y) I = restrict x I" "J \ I = {} \ restrict (merge I x J y) J = restrict y J" - by (auto simp: restrict_def intro!: ext) + by (auto simp: restrict_def) lemma extensional_insert_undefined[intro, simp]: assumes "a \ extensional (insert i I)" @@ -130,16 +130,16 @@ using assms by (auto simp: restrict_Pi_cancel) lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" - by (auto simp: restrict_def intro!: ext) + by (auto simp: restrict_def) lemma merge_restrict[simp]: "merge I (restrict x I) J y = merge I x J y" "merge I x J (restrict y J) = merge I x J y" - unfolding merge_def by (auto intro!: ext) + unfolding merge_def by auto lemma merge_x_x_eq_restrict[simp]: "merge I x J x = restrict x (I \ J)" - unfolding merge_def by (auto intro!: ext) + unfolding merge_def by auto lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" apply auto @@ -233,339 +233,355 @@ section "Products" -locale product_sigma_algebra = - fixes M :: "'i \ ('a, 'b) measure_space_scheme" - assumes sigma_algebras: "\i. sigma_algebra (M i)" +definition prod_emb where + "prod_emb I M K X = (\x. restrict x K) -` X \ (PIE i:I. space (M i))" + +lemma prod_emb_iff: + "f \ prod_emb I M K X \ f \ extensional I \ (restrict f K \ X) \ (\i\I. f i \ space (M i))" + unfolding prod_emb_def by auto -locale finite_product_sigma_algebra = product_sigma_algebra M - for M :: "'i \ ('a, 'b) measure_space_scheme" + - fixes I :: "'i set" - assumes finite_index[simp, intro]: "finite I" +lemma + shows prod_emb_empty[simp]: "prod_emb M L K {} = {}" + and prod_emb_Un[simp]: "prod_emb M L K (A \ B) = prod_emb M L K A \ prod_emb M L K B" + and prod_emb_Int: "prod_emb M L K (A \ B) = prod_emb M L K A \ prod_emb M L K B" + and prod_emb_UN[simp]: "prod_emb M L K (\i\I. F i) = (\i\I. prod_emb M L K (F i))" + and prod_emb_INT[simp]: "I \ {} \ prod_emb M L K (\i\I. F i) = (\i\I. prod_emb M L K (F i))" + and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B" + by (auto simp: prod_emb_def) -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)) \" +lemma prod_emb_PiE: "J \ I \ (\i. i \ J \ E i \ space (M i)) \ + prod_emb I M J (\\<^isub>E i\J. E i) = (\\<^isub>E i\I. if i \ J then E i else space (M i))" + by (force simp: prod_emb_def Pi_iff split_if_mem2) + +lemma prod_emb_PiE_same_index[simp]: "(\i. i \ I \ E i \ space (M i)) \ prod_emb I M I (Pi\<^isub>E I E) = Pi\<^isub>E I E" + by (auto simp: prod_emb_def Pi_iff) -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))))\" +definition PiM :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) measure" where + "PiM I M = extend_measure (\\<^isub>E i\I. space (M i)) + {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))} + (\(J, X). prod_emb I M J (\\<^isub>E j\J. X j)) + (\(J, X). \j\J \ {i\I. emeasure (M i) (space (M i)) \ 1}. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" + +definition prod_algebra :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) set set" where + "prod_algebra I M = (\(J, X). prod_emb I M J (\\<^isub>E j\J. X j)) ` + {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))}" + +abbreviation + "Pi\<^isub>M I M \ PiM I M" syntax - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10) + "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3PIM _:_./ _)" 10) syntax (xsymbols) - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) + "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^isub>M _\_./ _)" 10) syntax (HTML output) - "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => - ('i => 'a, 'b) measure_space_scheme" ("(3\\<^isub>M _\_./ _)" 10) + "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^isub>M _\_./ _)" 10) translations - "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" - -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 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 + "PIM x:I. M" == "CONST PiM I (%x. M)" -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 sigma_algebra_product_algebra . +lemma prod_algebra_sets_into_space: + "prod_algebra I M \ Pow (\\<^isub>E i\I. space (M i))" + using assms by (auto simp: prod_emb_def prod_algebra_def) -lemma product_algebraE: - 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_generator_def by auto - -lemma product_algebra_generatorI[intro]: - assumes "E \ (\ i\I. sets (M i))" - 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 prod_algebra_eq_finite: + assumes I: "finite I" + shows "prod_algebra I M = {(\\<^isub>E i\I. X i) |X. X \ (\ j\I. sets (M j))}" (is "?L = ?R") +proof (intro iffI set_eqI) + fix A assume "A \ ?L" + then obtain J E where J: "J \ {} \ I = {}" "finite J" "J \ I" "\i\J. E i \ sets (M i)" + and A: "A = prod_emb I M J (PIE j:J. E j)" + by (auto simp: prod_algebra_def) + let ?A = "\\<^isub>E i\I. if i \ J then E i else space (M i)" + have A: "A = ?A" + unfolding A using J by (intro prod_emb_PiE sets_into_space) auto + show "A \ ?R" unfolding A using J top + by (intro CollectI exI[of _ "\i. if i \ J then E i else space (M i)"]) simp +next + fix A assume "A \ ?R" + then obtain X where "A = (\\<^isub>E i\I. X i)" and X: "X \ (\ j\I. sets (M j))" by auto + then have A: "A = prod_emb I M I (\\<^isub>E i\I. X i)" + using sets_into_space by (force simp: prod_emb_def Pi_iff) + from X I show "A \ ?L" unfolding A + by (auto simp: prod_algebra_def) +qed -lemma space_product_algebra[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 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 prod_algebraI: + "finite J \ (J \ {} \ I = {}) \ J \ I \ (\i. i \ J \ E i \ sets (M i)) + \ prod_emb I M J (PIE j:J. E j) \ prod_algebra I M" + by (auto simp: prod_algebra_def Pi_iff) -lemma product_algebra_generator_sets_into_space: - assumes "\i. i\I \ sets (M i) \ Pow (space (M i))" - 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: sets_product_algebra) - -lemma Int_stable_product_algebra_generator: - "(\i. i \ I \ Int_stable (M i)) \ Int_stable (product_algebra_generator I M)" - by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff) +lemma prod_algebraE: + assumes A: "A \ prod_algebra I M" + obtains J E where "A = prod_emb I M J (PIE j:J. E j)" + "finite J" "J \ {} \ I = {}" "J \ I" "\i. i \ J \ E i \ sets (M i)" + using A by (auto simp: prod_algebra_def) -section "Generating set generates also product algebra" +lemma prod_algebraE_all: + assumes A: "A \ prod_algebra I M" + obtains E where "A = Pi\<^isub>E I E" "E \ (\ i\I. sets (M i))" +proof - + from A obtain E J where A: "A = prod_emb I M J (Pi\<^isub>E J E)" + and J: "J \ I" and E: "E \ (\ i\J. sets (M i))" + by (auto simp: prod_algebra_def) + from E have "\i. i \ J \ E i \ space (M i)" + using sets_into_space by auto + then have "A = (\\<^isub>E i\I. if i\J then E i else space (M i))" + using A J by (auto simp: prod_emb_PiE) + moreover then have "(\i. if i\J then E i else space (M i)) \ (\ i\I. sets (M i))" + using top E by auto + ultimately show ?thesis using that by auto +qed -lemma sigma_product_algebra_sigma_eq: - assumes "finite I" - assumes mono: "\i. i \ I \ incseq (S i)" - assumes union: "\i. i \ I \ (\j. S i j) = 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 "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 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 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 mono union unfolding incseq_Suc_iff 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 space_product_algebra - apply simp - apply (subst Pi_UN[OF `finite I`]) - using mono[THEN incseqD] apply simp - apply (simp add: PiE_Int) - apply (intro PiE_cong) - using A sets_into by (auto intro!: into_space) - 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: 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 - have "Pi\<^isub>E I A = (\i\I. (\x. x i) -` (A i) \ space ?E)" - 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_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 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 sets_product_algebra - proof induct - case (Basic A) then show ?case - 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 +lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)" +proof (unfold Int_stable_def, safe) + fix A assume "A \ prod_algebra I M" + from prod_algebraE[OF this] guess J E . note A = this + fix B assume "B \ prod_algebra I M" + from prod_algebraE[OF this] guess K F . note B = this + have "A \ B = prod_emb I M (J \ K) (\\<^isub>E i\J \ K. (if i \ J then E i else space (M i)) \ + (if i \ K then F i else space (M i)))" + unfolding A B using A(2,3,4) A(5)[THEN sets_into_space] B(2,3,4) B(5)[THEN sets_into_space] + apply (subst (1 2 3) prod_emb_PiE) + apply (simp_all add: subset_eq PiE_Int) + apply blast + apply (intro PiE_cong) + apply auto + done + also have "\ \ prod_algebra I M" + using A B by (auto intro!: prod_algebraI) + finally show "A \ B \ prod_algebra I M" . +qed + +lemma prod_algebra_mono: + assumes space: "\i. i \ I \ space (E i) = space (F i)" + assumes sets: "\i. i \ I \ sets (E i) \ sets (F i)" + shows "prod_algebra I E \ prod_algebra I F" +proof + fix A assume "A \ prod_algebra I E" + then obtain J G where J: "J \ {} \ I = {}" "finite J" "J \ I" + and A: "A = prod_emb I E J (\\<^isub>E i\J. G i)" + and G: "\i. i \ J \ G i \ sets (E i)" + by (auto simp: prod_algebra_def) + moreover + from space have "(\\<^isub>E i\I. space (E i)) = (\\<^isub>E i\I. space (F i))" + by (rule PiE_cong) + with A have "A = prod_emb I F J (\\<^isub>E i\J. G i)" + by (simp add: prod_emb_def) + moreover + from sets G J have "\i. i \ J \ G i \ sets (F i)" + by auto + ultimately show "A \ prod_algebra I F" + apply (simp add: prod_algebra_def image_iff) + apply (intro exI[of _ J] exI[of _ G] conjI) + apply auto + done 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 space_PiM: "space (\\<^isub>M i\I. M i) = (\\<^isub>E i\I. space (M i))" + using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp + +lemma sets_PiM: "sets (\\<^isub>M i\I. M i) = sigma_sets (\\<^isub>E i\I. space (M i)) (prod_algebra I M)" + using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp -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) +lemma sets_PiM_single: "sets (PiM I M) = + sigma_sets (\\<^isub>E i\I. space (M i)) {{f\\\<^isub>E i\I. space (M i). f i \ A} | i A. i \ I \ A \ sets (M i)}" + (is "_ = sigma_sets ?\ ?R") + unfolding sets_PiM +proof (rule sigma_sets_eqI) + interpret R: sigma_algebra ?\ "sigma_sets ?\ ?R" by (rule sigma_algebra_sigma_sets) auto + fix A assume "A \ prod_algebra I M" + from prod_algebraE[OF this] guess J X . note X = this + show "A \ sigma_sets ?\ ?R" + proof cases + assume "I = {}" + with X have "A = {\x. undefined}" by (auto simp: prod_emb_def) + with `I = {}` show ?thesis by (auto intro!: sigma_sets_top) + next + assume "I \ {}" + with X have "A = (\j\J. {f\(\\<^isub>E i\I. space (M i)). f j \ X j})" + using sets_into_space[OF X(5)] + by (auto simp: prod_emb_PiE[OF _ sets_into_space] Pi_iff split: split_if_asm) blast + also have "\ \ sigma_sets ?\ ?R" + using X `I \ {}` by (intro R.finite_INT sigma_sets.Basic) auto + finally show "A \ sigma_sets ?\ ?R" . + qed +next + fix A assume "A \ ?R" + then obtain i B where A: "A = {f\\\<^isub>E i\I. space (M i). f i \ B}" "i \ I" "B \ sets (M i)" + by auto + then have "A = prod_emb I M {i} (\\<^isub>E i\{i}. B)" + using sets_into_space[OF A(3)] + apply (subst prod_emb_PiE) + apply (auto simp: Pi_iff split: split_if_asm) + apply blast + done + also have "\ \ sigma_sets ?\ (prod_algebra I M)" + using A by (intro sigma_sets.Basic prod_algebraI) auto + finally show "A \ sigma_sets ?\ (prod_algebra I M)" . +qed + +lemma sets_PiM_I: + assumes "finite J" "J \ I" "\i\J. E i \ sets (M i)" + shows "prod_emb I M J (PIE j:J. E j) \ sets (PIM i:I. M i)" +proof cases + assume "J = {}" + then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))" + by (auto simp: prod_emb_def) + then show ?thesis + by (auto simp add: sets_PiM intro!: sigma_sets_top) +next + assume "J \ {}" with assms show ?thesis + by (auto simp add: sets_PiM prod_algebra_def intro!: sigma_sets.Basic) qed -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) +lemma measurable_PiM: + assumes space: "f \ space N \ (\\<^isub>E i\I. space (M i))" + assumes sets: "\X J. J \ {} \ I = {} \ finite J \ J \ I \ (\i. i \ J \ X i \ sets (M i)) \ + f -` prod_emb I M J (Pi\<^isub>E J X) \ space N \ sets N" + shows "f \ measurable N (PiM I M)" + using sets_PiM prod_algebra_sets_into_space space +proof (rule measurable_sigma_sets) + fix A assume "A \ prod_algebra I M" + from prod_algebraE[OF this] guess J X . + with sets[of J X] show "f -` A \ space N \ sets N" by auto +qed + +lemma measurable_PiM_Collect: + assumes space: "f \ space N \ (\\<^isub>E i\I. space (M i))" + assumes sets: "\X J. J \ {} \ I = {} \ finite J \ J \ I \ (\i. i \ J \ X i \ sets (M i)) \ + {\\space N. \i\J. f \ i \ X i} \ sets N" + shows "f \ measurable N (PiM I M)" + using sets_PiM prod_algebra_sets_into_space space +proof (rule measurable_sigma_sets) + fix A assume "A \ prod_algebra I M" + from prod_algebraE[OF this] guess J X . note X = this + have "f -` A \ space N = {\ \ space N. \i\J. f \ i \ X i}" + using sets_into_space[OF X(5)] X(2-) space unfolding X(1) + by (subst prod_emb_PiE) (auto simp: Pi_iff split: split_if_asm) + also have "\ \ sets N" using X(3,2,4,5) by (rule sets) + finally show "f -` A \ space N \ sets N" . qed -lemma (in product_sigma_algebra) measurable_merge: - assumes [simp]: "I \ J = {}" - 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 - - 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) fast+ - 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 measurable_PiM_single: + assumes space: "f \ space N \ (\\<^isub>E i\I. space (M i))" + assumes sets: "\A i. i \ I \ A \ sets (M i) \ {\ \ space N. f \ i \ A} \ sets N" + shows "f \ measurable N (PiM I M)" + using sets_PiM_single +proof (rule measurable_sigma_sets) + fix A assume "A \ {{f \ \\<^isub>E i\I. space (M i). f i \ A} |i A. i \ I \ A \ sets (M i)}" + then obtain B i where "A = {f \ \\<^isub>E i\I. space (M i). f i \ B}" and B: "i \ I" "B \ sets (M i)" + by auto + with space have "f -` A \ space N = {\ \ space N. f \ i \ B}" by auto + also have "\ \ sets N" using B by (rule sets) + finally show "f -` A \ space N \ sets N" . +qed (auto simp: space) -lemma (in product_sigma_algebra) measurable_component_singleton: +lemma sets_PiM_I_finite[simp, intro]: + assumes "finite I" and sets: "(\i. i \ I \ E i \ sets (M i))" + shows "(PIE j:I. E j) \ sets (PIM i:I. M i)" + using sets_PiM_I[of I I E M] sets_into_space[OF sets] `finite I` sets by auto + +lemma 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 \ _") +proof (intro measurable_PiM_single) + fix j A assume "j \ insert i I" "A \ sets (M j)" + moreover have "{\ \ space (M i). (x(i := \)) j \ A} = + (if i = j then space (M i) \ A else if x j \ A then space (M i) else {})" + by auto + ultimately show "{\ \ space (M i). (x(i := \)) j \ A} \ sets (M i)" + by auto +qed (insert sets_into_space assms, auto simp: space_PiM) + +lemma measurable_component_singleton: 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 (fastforce dest: Pi_mem split: split_if_asm) + then have "(\x. x i) -` A \ space (Pi\<^isub>M I M) = prod_emb I M {i} (\\<^isub>E j\{i}. A)" + using sets_into_space `i \ I` + by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM 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) + using `A \ sets (M i)` `i \ I` by (auto intro!: sets_PiM_I) +qed (insert `i \ I`, auto simp: space_PiM) + +lemma 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)" + (is "?f \ measurable ?P ?I") +proof (rule measurable_PiM_single) + fix j A assume j: "j \ insert i I" and A: "A \ sets (M j)" + have "{\ \ space ?P. (\(f, x). fun_upd f i x) \ j \ A} = + (if j = i then space (Pi\<^isub>M I M) \ A else ((\x. x j) \ fst) -` A \ space ?P)" + using sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM) + also have "\ \ sets ?P" + using A j + by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) + finally show "{\ \ space ?P. prod_case (\f. fun_upd f i) \ j \ A} \ sets ?P" . +qed (auto simp: space_pair_measure space_PiM) -lemma (in sigma_algebra) measurable_restrict: - assumes I: "finite I" - assumes "\i. i \ I \ sets (N i) \ Pow (space (N i))" - assumes X: "\i. i \ I \ X i \ measurable M (N i)" - shows "(\x. \i\I. X i x) \ measurable M (Pi\<^isub>M I N)" - unfolding product_algebra_def -proof (simp, rule measurable_sigma) - show "sets (product_algebra_generator I N) \ Pow (space (product_algebra_generator I N))" - by (rule product_algebra_generator_sets_into_space) fact - show "(\x. \i\I. X i x) \ space M \ space (product_algebra_generator I N)" - using X by (auto simp: measurable_def) - fix E assume "E \ sets (product_algebra_generator I N)" - then obtain F where "E = Pi\<^isub>E I F" and F: "\i. i \ I \ F i \ sets (N i)" - by (auto simp: product_algebra_generator_def) - then have "(\x. \i\I. X i x) -` E \ space M = (\i\I. X i -` F i \ space M) \ space M" - by (auto simp: Pi_iff) - also have "\ \ sets M" - proof cases - assume "I = {}" then show ?thesis by simp - next - assume "I \ {}" with X F I show ?thesis - by (intro finite_INT measurable_sets Int) auto - qed - finally show "(\x. \i\I. X i x) -` E \ space M \ sets M" . -qed +lemma measurable_merge: + assumes "I \ J = {}" + 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)" + (is "?f \ measurable ?P ?U") +proof (rule measurable_PiM_single) + fix i A assume A: "A \ sets (M i)" "i \ I \ J" + then have "{\ \ space ?P. prod_case (\x. merge I x J) \ i \ A} = + (if i \ I then ((\x. x i) \ fst) -` A \ space ?P else ((\x. x i) \ snd) -` A \ space ?P)" + using `I \ J = {}` by auto + also have "\ \ sets ?P" + using A + by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) + finally show "{\ \ space ?P. prod_case (\x. merge I x J) \ i \ A} \ sets ?P" . +qed (insert assms, auto simp: space_pair_measure space_PiM) -locale product_sigma_finite = product_sigma_algebra M - for M :: "'i \ ('a,'b) measure_space_scheme" + +lemma measurable_restrict: + assumes X: "\i. i \ I \ X i \ measurable N (M i)" + shows "(\x. \i\I. X i x) \ measurable N (Pi\<^isub>M I M)" +proof (rule measurable_PiM_single) + fix A i assume A: "i \ I" "A \ sets (M i)" + then have "{\ \ space N. (\i\I. X i \) i \ A} = X i -` A \ space N" + by auto + then show "{\ \ space N. (\i\I. X i \) i \ A} \ sets N" + using A X by (auto intro!: measurable_sets) +qed (insert X, auto dest: measurable_space) + +locale product_sigma_finite = + fixes M :: "'i \ 'a measure" assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" sublocale product_sigma_finite \ M: sigma_finite_measure "M i" for i by (rule sigma_finite_measures) -locale finite_product_sigma_finite = finite_product_sigma_algebra M I + product_sigma_finite M - for M :: "'i \ ('a,'b) measure_space_scheme" and I :: "'i set" - -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_ereal_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_ereal_0 intro!: bexI) - ultimately show ?thesis - unfolding product_algebra_generator_def by simp -qed +locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \ 'a measure" + + fixes I :: "'i set" + assumes finite_index: "finite I" lemma (in finite_product_sigma_finite) sigma_finite_pairs: "\F::'i \ nat \ 'a set. (\i\I. range (F i) \ sets (M i)) \ - (\k. \i\I. \ i (F i k) \ \) \ incseq (\k. \\<^isub>E i\I. F i k) \ - (\k. \\<^isub>E i\I. F i k) = space G" + (\k. \i\I. emeasure (M i) (F i k) \ \) \ incseq (\k. \\<^isub>E i\I. F i k) \ + (\k. \\<^isub>E i\I. F i k) = space (PiM I M)" proof - - have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. \ i (F k) \ \)" - using M.sigma_finite_up by simp + have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. emeasure (M i) (F k) \ \)" + using M.sigma_finite_incseq by metis from choice[OF this] guess F :: "'i \ nat \ 'a set" .. - then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. \ i (F i k) \ \" + then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. emeasure (M i) (F i k) \ \" by auto let ?F = "\k. \\<^isub>E i\I. F i k" - note space_product_algebra[simp] + note space_PiM[simp] show ?thesis proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) fix i show "range (F i) \ sets (M i)" by fact next - fix i k show "\ i (F i k) \ \" by fact + fix i k show "emeasure (M i) (F i k) \ \" by fact next - fix A assume "A \ (\i. ?F i)" then show "A \ space G" - using `\i. range (F i) \ sets (M i)` M.sets_into_space - by (force simp: image_subset_iff) + fix A assume "A \ (\i. ?F i)" then show "A \ space (PiM I M)" + using `\i. range (F i) \ sets (M i)` sets_into_space + by auto blast next - fix f assume "f \ space G" + fix f assume "f \ space (PiM I M)" with Pi_UN[OF finite_index, of "\k i. F i k"] F show "f \ (\i. ?F i)" by (auto simp: incseq_def) next @@ -574,140 +590,144 @@ qed qed -lemma sets_pair_cancel_measure[simp]: - "sets (M1\measure := m1\ \\<^isub>M M2) = sets (M1 \\<^isub>M M2)" - "sets (M1 \\<^isub>M M2\measure := m2\) = sets (M1 \\<^isub>M M2)" - unfolding pair_measure_def pair_measure_generator_def sets_sigma - by simp_all - -lemma measurable_pair_cancel_measure[simp]: - "measurable (M1\measure := m1\ \\<^isub>M M2) M = measurable (M1 \\<^isub>M M2) M" - "measurable (M1 \\<^isub>M M2\measure := m2\) M = measurable (M1 \\<^isub>M M2) M" - "measurable M (M1\measure := m3\ \\<^isub>M M2) = measurable M (M1 \\<^isub>M M2)" - "measurable M (M1 \\<^isub>M M2\measure := m4\) = measurable M (M1 \\<^isub>M M2)" - unfolding measurable_def by (auto simp add: space_pair_measure) - -lemma (in product_sigma_finite) product_measure_exists: +lemma (in product_sigma_finite) assumes "finite 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)))" + shows sigma_finite: "sigma_finite_measure (PiM I M)" + and emeasure_PiM: + "\A. (\i. i\I \ A i \ sets (M i)) \ emeasure (PiM I M) (Pi\<^isub>E I A) = (\i\I. emeasure (M i) (A i))" using `finite I` proof induct case empty - interpret finite_product_sigma_finite M "{}" by default simp - let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ ereal" - 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 + let ?\ = "\A. if A = {} then 0 else (1::ereal)" + have "prod_algebra {} M = {{\_. undefined}}" + by (auto simp: prod_algebra_def prod_emb_def intro!: image_eqI) + then have sets_empty: "sets (PiM {} M) = {{}, {\_. undefined}}" + by (simp add: sets_PiM) + have "emeasure (Pi\<^isub>M {} M) (prod_emb {} M {} (\\<^isub>E i\{}. {})) = 1" + proof (subst emeasure_extend_measure_Pair[OF PiM_def]) + have "finite (space (PiM {} M))" + by (simp add: space_PiM) + moreover show "positive (PiM {} M) ?\" + by (auto simp: positive_def) + ultimately show "countably_additive (PiM {} M) ?\" + by (rule countably_additiveI_finite) (auto simp: additive_def space_PiM sets_empty) + qed (auto simp: prod_emb_def) + also have *: "(prod_emb {} M {} (\\<^isub>E i\{}. {})) = {\_. undefined}" + by (auto simp: prod_emb_def) + finally have emeasure_eq: "emeasure (Pi\<^isub>M {} M) {\_. undefined} = 1" . + + interpret finite_measure "PiM {} M" + by default (simp add: space_PiM emeasure_eq) + case 1 show ?case .. + + case 2 show ?case + using emeasure_eq * by simp next case (insert i I) 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 - from insert obtain \ where - 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" .. + interpret I: sigma_finite_measure "PiM I M" by fact + interpret P: pair_sigma_finite "PiM I M" "M i" .. let ?h = "(\(f, y). f(i := y))" - 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']) (auto simp add: measure_preserving_def) - 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: 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 fastforce - using A apply fastforce - using `i \ I` `finite I` prod[of A] A by (auto simp: ac_simps) } - note product = this - 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)" - "incseq (\k. \\<^isub>E j \ insert i I. F j k)" - "(\k. \\<^isub>E j \ insert i I. F j k) = space I'.G" - "\k. \j. j \ insert i I \ \ j (F j k) \ \" - 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) = (\\<^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(3) show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp - next - fix j - have "\k. k \ insert i I \ 0 \ measure (M k) (F k j)" - using F(1) by auto - with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\ (?F j) \ \" - by (subst product) auto - qed - qed + + let ?P = "distr (Pi\<^isub>M I M \\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h" + let ?\ = "emeasure ?P" + let ?I = "{j \ insert i I. emeasure (M j) (space (M j)) \ 1}" + let ?f = "\J E j. if j \ J then emeasure (M j) (E j) else emeasure (M j) (space (M j))" + + { case 2 + have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>E (insert i I) A)) = + (\i\insert i I. emeasure (M i) (A i))" + proof (subst emeasure_extend_measure_Pair[OF PiM_def]) + fix J E assume "(J \ {} \ insert i I = {}) \ finite J \ J \ insert i I \ E \ (\ j\J. sets (M j))" + then have J: "J \ {}" "finite J" "J \ insert i I" and E: "\j\J. E j \ sets (M j)" by auto + let ?p = "prod_emb (insert i I) M J (Pi\<^isub>E J E)" + let ?p' = "prod_emb I M (J - {i}) (\\<^isub>E j\J-{i}. E j)" + have "?\ ?p = + emeasure (Pi\<^isub>M I M \\<^isub>M (M i)) (?h -` ?p \ space (Pi\<^isub>M I M \\<^isub>M M i))" + by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+ + also have "?h -` ?p \ space (Pi\<^isub>M I M \\<^isub>M M i) = ?p' \ (if i \ J then E i else space (M i))" + using J E[rule_format, THEN sets_into_space] + by (force simp: space_pair_measure space_PiM Pi_iff prod_emb_iff split: split_if_asm) + also have "emeasure (Pi\<^isub>M I M \\<^isub>M (M i)) (?p' \ (if i \ J then E i else space (M i))) = + emeasure (Pi\<^isub>M I M) ?p' * emeasure (M i) (if i \ J then (E i) else space (M i))" + using J E by (intro P.emeasure_pair_measure_Times sets_PiM_I) auto + also have "?p' = (\\<^isub>E j\I. if j \ J-{i} then E j else space (M j))" + using J E[rule_format, THEN sets_into_space] + by (auto simp: prod_emb_iff Pi_iff split: split_if_asm) blast+ + also have "emeasure (Pi\<^isub>M I M) (\\<^isub>E j\I. if j \ J-{i} then E j else space (M j)) = + (\ j\I. if j \ J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))" + using E by (subst insert) (auto intro!: setprod_cong) + also have "(\j\I. if j \ J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * + emeasure (M i) (if i \ J then E i else space (M i)) = (\j\insert i I. ?f J E j)" + using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod_cong) + also have "\ = (\j\J \ ?I. ?f J E j)" + using insert(1,2) J E by (intro setprod_mono_one_right) auto + finally show "?\ ?p = \" . + + show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \ Pow (\\<^isub>E i\insert i I. space (M i))" + using J E[rule_format, THEN sets_into_space] by (auto simp: prod_emb_iff) + next + show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\" "countably_additive (sets (Pi\<^isub>M (insert i I) M)) ?\" + using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all + next + show "(insert i I \ {} \ insert i I = {}) \ finite (insert i I) \ + insert i I \ insert i I \ A \ (\ j\insert i I. sets (M j))" + using insert(1,2) 2 by auto + qed (auto intro!: setprod_cong) + with 2[THEN sets_into_space] show ?case by (subst (asm) prod_emb_PiE_same_index) auto } + note product = this + + 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)" + "incseq (\k. \\<^isub>E j \ insert i I. F j k)" + "(\k. \\<^isub>E j \ insert i I. F j k) = space (Pi\<^isub>M (insert i I) M)" + "\k. \j. j \ insert i I \ emeasure (M j) (F j k) \ \" + by blast+ + let ?F = "\k. \\<^isub>E j \ insert i I. F j k" + + case 1 show ?case + proof (unfold_locales, intro exI[of _ ?F] conjI allI) + show "range ?F \ sets (Pi\<^isub>M (insert i I) M)" using F(1) insert(1,2) by auto + next + from F(3) show "(\i. ?F i) = space (Pi\<^isub>M (insert i I) M)" by simp + next + fix j + from F `finite I` setprod_PInf[of "insert i I", OF emeasure_nonneg, of M] + show "emeasure (\\<^isub>M i\insert i I. M i) (?F j) \ \" + by (subst product) auto qed 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 +sublocale finite_product_sigma_finite \ sigma_finite_measure "Pi\<^isub>M I M" + using sigma_finite[OF finite_index] . lemma (in finite_product_sigma_finite) measure_times: - assumes "\i. i \ I \ A i \ sets (M 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. 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 + "(\i. i \ I \ A i \ sets (M i)) \ emeasure (Pi\<^isub>M I M) (Pi\<^isub>E I A) = (\i\I. emeasure (M i) (A i))" + using emeasure_PiM[OF finite_index] by auto lemma (in product_sigma_finite) product_measure_empty[simp]: - "measure (Pi\<^isub>M {} M) {\x. undefined} = 1" + "emeasure (Pi\<^isub>M {} M) {\x. undefined} = 1" proof - 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 + shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\k. undefined}" + and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\k. undefined} }" + by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq) lemma (in product_sigma_finite) positive_integral_empty: assumes pos: "0 \ f (\k. undefined)" shows "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>M {} M) (Pi\<^isub>E {} A) = 1" + have "\A. emeasure (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 [abs_def] - proof (simp add: P_empty, intro antisym) + unfolding positive_integral_def simple_function_def simple_integral_def[abs_def] + proof (simp add: space_PiM_empty sets_PiM_empty, intro antisym) show "f (\k. undefined) \ (SUP f:{g. g \ max 0 \ f}. f (\k. undefined))" by (intro SUP_upper) (auto simp: le_fun_def split: split_max) show "(SUP f:{g. g \ max 0 \ f}. f (\k. undefined)) \ f (\k. undefined)" using pos @@ -715,71 +735,72 @@ qed qed -lemma (in product_sigma_finite) measure_fold: +lemma (in product_sigma_finite) distr_merge: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - assumes A: "A \ sets (Pi\<^isub>M (I \ J) M)" - shows "measure (Pi\<^isub>M (I \ J) M) A = - 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))" + shows "distr (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M) (\(x,y). merge I x J y) = Pi\<^isub>M (I \ J) M" + (is "?D = ?P") proof - 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 J.P by default + interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" 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 F: "\i. i\ I \ J \ range (F i) \ sets (M i)" "incseq (\k. \\<^isub>E i\I \ J. F i k)" - "(\k. \\<^isub>E i\I \ J. F i k) = space IJ.G" - "\k. \i\I\J. \ i (F i k) \ \" + "(\k. \\<^isub>E i\I \ J. F i k) = space ?P" + "\k. \i\I\J. emeasure (M i) (F i k) \ \" by auto let ?F = "\k. \\<^isub>E i\I \ J. F i k" - show "IJ.\ A = P.\ (?X A)" - proof (rule measure_unique_Int_stable_vimage) - show "measure_space IJ.P" "measure_space P.P" by default - show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \ sets (sigma IJ.G)" - using A unfolding product_algebra_def by auto - next - show "Int_stable IJ.G" - by (rule Int_stable_product_algebra_generator) - (auto simp: Int_stable_def) - show "range ?F \ sets IJ.G" using F - by (simp add: image_subset_iff product_algebra_def - product_algebra_generator_def) - show "incseq ?F" "(\i. ?F i) = space IJ.G " by fact+ + + show ?thesis + proof (rule measure_eqI_generator_eq[symmetric]) + show "Int_stable (prod_algebra (I \ J) M)" + by (rule Int_stable_prod_algebra) + show "prod_algebra (I \ J) M \ Pow (\\<^isub>E i \ I \ J. space (M i))" + by (rule prod_algebra_sets_into_space) + show "sets ?P = sigma_sets (\\<^isub>E i\I \ J. space (M i)) (prod_algebra (I \ J) M)" + by (rule sets_PiM) + then show "sets ?D = sigma_sets (\\<^isub>E i\I \ J. space (M i)) (prod_algebra (I \ J) M)" + by simp + + show "range ?F \ prod_algebra (I \ J) M" using F + using fin by (auto simp: prod_algebra_eq_finite) + show "incseq ?F" by fact + show "(\i. \\<^isub>E ia\I \ J. F ia i) = (\\<^isub>E i\I \ J. space (M i))" + using F(3) by (simp add: space_PiM) next fix k - have "\j. j \ I \ J \ 0 \ measure (M j) (F j k)" - using F(1) by auto - with F `finite I` setprod_PInf[of "I \ J", OF this] - show "IJ.\ (?F k) \ \" by (subst IJ.measure_times) auto + from F `finite I` setprod_PInf[of "I \ J", OF emeasure_nonneg, of M] + show "emeasure ?P (?F k) \ \" by (subst IJ.measure_times) auto next - fix A assume "A \ sets IJ.G" - then obtain F where A: "A = Pi\<^isub>E (I \ J) F" - and F: "\i. i \ I \ J \ F i \ sets (M i)" - by (auto simp: product_algebra_generator_def) - then have X: "?X A = (Pi\<^isub>E I F \ Pi\<^isub>E J F)" - 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))" + fix A assume A: "A \ prod_algebra (I \ J) M" + with fin obtain F where A_eq: "A = (Pi\<^isub>E (I \ J) F)" and F: "\i\I \ J. F i \ sets (M i)" + by (auto simp add: prod_algebra_eq_finite) + let ?B = "Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M" + let ?X = "?g -` A \ space ?B" + have "Pi\<^isub>E I F \ space (Pi\<^isub>M I M)" "Pi\<^isub>E J F \ space (Pi\<^isub>M J M)" + using F[rule_format, THEN sets_into_space] by (auto simp: space_PiM) + then have X: "?X = (Pi\<^isub>E I F \ Pi\<^isub>E J F)" + unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM) + have "emeasure ?D A = emeasure ?B ?X" + using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM) + also have "emeasure ?B ?X = (\i\I. emeasure (M i) (F i)) * (\i\J. emeasure (M i) (F i))" + using `finite J` `finite I` F X + by (simp add: P.emeasure_pair_measure_Times I.measure_times J.measure_times Pi_iff) + also have "\ = (\i\I \ J. emeasure (M i) (F i))" using `finite J` `finite I` `I \ J = {}` by (simp add: setprod_Un_one) - also have "\ = IJ.\ A" + also have "\ = emeasure ?P (Pi\<^isub>E (I \ J) F)" using `finite J` `finite I` F unfolding A by (intro IJ.measure_times[symmetric]) auto - finally show "IJ.\ A = P.\ (?X A)" by (rule sym) - qed (rule measurable_merge[OF IJ]) + finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp + qed qed -lemma (in product_sigma_finite) measure_preserving_merge: - assumes IJ: "I \ J = {}" and "finite I" "finite J" - shows "(\(x,y). merge I x J y) \ measure_preserving (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M)" - by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms) - lemma (in product_sigma_finite) product_positive_integral_fold: - assumes IJ[simp]: "I \ J = {}" "finite I" "finite J" + assumes IJ: "I \ J = {}" "finite I" "finite J" 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))" @@ -787,42 +808,38 @@ interpret I: finite_product_sigma_finite M I by default fact interpret J: finite_product_sigma_finite M J by default fact interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default - interpret IJ: finite_product_sigma_finite M "I \ J" by default simp - have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable P.P" + have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)" using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def) show ?thesis - unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] - proof (rule P.positive_integral_vimage) - show "sigma_algebra IJ.P" by default - show "(\(x, y). merge I x J y) \ measure_preserving P.P IJ.P" - using IJ by (rule measure_preserving_merge) - show "f \ borel_measurable IJ.P" using f by simp - qed + apply (subst distr_merge[OF IJ, symmetric]) + apply (subst positive_integral_distr[OF measurable_merge f, OF IJ(1)]) + apply (subst P.positive_integral_fst_measurable(2)[symmetric, OF P_borel]) + apply simp + done qed -lemma (in product_sigma_finite) measure_preserving_component_singelton: - "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" -proof (intro measure_preservingI measurable_component_singleton) +lemma (in product_sigma_finite) distr_singleton: + "distr (Pi\<^isub>M {i} M) (M i) (\x. x i) = M i" (is "?D = _") +proof (intro measure_eqI[symmetric]) interpret I: finite_product_sigma_finite M "{i}" by default simp - 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 "I.\ ?P = M.\ i A" unfolding * - using A I.measure_times[of "\_. A"] by auto -qed auto + fix A assume A: "A \ sets (M i)" + moreover then have "(\x. x i) -` A \ space (Pi\<^isub>M {i} M) = (\\<^isub>E i\{i}. A)" + using sets_into_space by (auto simp: space_PiM) + ultimately show "emeasure (M i) A = emeasure ?D A" + using A I.measure_times[of "\_. A"] + by (simp add: emeasure_distr measurable_component_singleton) +qed simp lemma (in product_sigma_finite) product_positive_integral_singleton: assumes f: "f \ borel_measurable (M i)" 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 - show ?thesis - proof (rule I.positive_integral_vimage[symmetric]) - show "sigma_algebra (M i)" by (rule sigma_algebras) - show "(\x. x i) \ measure_preserving (Pi\<^isub>M {i} M) (M i)" - by (rule measure_preserving_component_singelton) - show "f \ borel_measurable (M i)" by fact - qed + from f show ?thesis + apply (subst distr_singleton[symmetric]) + apply (subst positive_integral_distr[OF measurable_component_singleton]) + apply simp_all + done qed lemma (in product_sigma_finite) product_positive_integral_insert: @@ -832,19 +849,18 @@ proof - 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} = {}" and insert: "I \ {i} = insert i I" using f by auto show ?thesis 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" + proof (rule positive_integral_cong, subst product_positive_integral_singleton) + fix x assume x: "x \ space (Pi\<^isub>M I M)" 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) + using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def space_PiM) 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) + using measurable_comp[OF measurable_component_update f, OF x `i \ I`] + unfolding comp_def . show "integral\<^isup>P (M i) ?f = \\<^isup>+ y. f (x(i:=y)) \M i" unfolding f'_eq by simp qed @@ -856,10 +872,6 @@ and pos: "\i x. i \ I \ 0 \ f i x" 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 - 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 @@ -867,16 +879,16 @@ using insert by (auto intro!: setprod_cong) 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_ereal_setprod sigma_algebra_product_algebra + by (intro borel_measurable_ereal_setprod measurable_comp[OF measurable_component_singleton, unfolded comp_def]) auto then show ?case apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) - apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc) - apply (subst I.positive_integral_cmult) - apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive) + apply (simp add: insert(2-) * pos borel setprod_ereal_pos positive_integral_multc) + apply (subst positive_integral_cmult) + apply (auto simp add: pos borel insert(2-) setprod_ereal_pos positive_integral_positive) done -qed +qed (simp add: space_PiM) lemma (in product_sigma_finite) product_integral_singleton: assumes f: "f \ borel_measurable (M i)" @@ -899,48 +911,44 @@ 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 J.P by default + interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default let ?M = "\(x, y). merge I x J y" let ?f = "\x. f (?M x)" + from f have f_borel: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" + by auto + have P_borel: "(\x. f (case x of (x, y) \ merge I x J y)) \ borel_measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)" + using measurable_comp[OF measurable_merge[OF IJ(1)] f_borel] by (simp add: comp_def) + have f_int: "integrable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) ?f" + by (rule integrable_distr[OF measurable_merge[OF IJ]]) (simp add: distr_merge[OF IJ fin] f) show ?thesis - proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) - have 1: "sigma_algebra IJ.P" by default - have 2: "?M \ measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] . - have 3: "integrable (Pi\<^isub>M (I \ J) M) f" by fact - then have 4: "f \ borel_measurable (Pi\<^isub>M (I \ J) M)" - by (simp add: integrable_def) - show "integrable P.P ?f" - by (rule P.integrable_vimage[where f=f, OF 1 2 3]) - show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" - by (rule P.integral_vimage[where f=f, OF 1 2 4]) - qed + apply (subst distr_merge[symmetric, OF IJ fin]) + apply (subst integral_distr[OF measurable_merge[OF IJ] f_borel]) + apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int]) + apply simp + done qed lemma (in product_sigma_finite) product_integral_insert: - assumes [simp]: "finite I" "i \ I" + assumes I: "finite I" "i \ I" 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_finite I.P i.P .. - have IJ: "I \ {i} = {}" by auto - show ?thesis - unfolding product_integral_fold[OF IJ, simplified, OF f] - proof (rule I.integral_cong, subst product_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) - have f: "f \ borel_measurable I'.P" using f unfolding integrable_def by auto - show "?f \ borel_measurable (M i)" - 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 + have "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = integral\<^isup>L (Pi\<^isub>M (I \ {i}) M) f" + by simp + also have "\ = (\x. (\y. f (merge I x {i} y) \Pi\<^isub>M {i} M) \Pi\<^isub>M I M)" + using f I by (intro product_integral_fold) auto + also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^isub>M I M)" + proof (rule integral_cong, subst product_integral_singleton[symmetric]) + fix x assume x: "x \ space (Pi\<^isub>M I M)" + have f_borel: "f \ borel_measurable (Pi\<^isub>M (insert i I) M)" + using f by auto + show "(\y. f (x(i := y))) \ borel_measurable (M i)" + using measurable_comp[OF measurable_component_update f_borel, OF x `i \ I`] + unfolding comp_def . + from x I show "(\ y. f (merge I x {i} y) \Pi\<^isub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^isub>M {i} M)" + by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def) qed + finally show ?thesis . qed lemma (in product_sigma_finite) product_integrable_setprod: @@ -951,24 +959,23 @@ interpret finite_product_sigma_finite M I by default fact have f: "\i. i \ I \ f i \ borel_measurable (M i)" using integrable unfolding integrable_def by auto - then have borel: "?f \ borel_measurable P" - using measurable_comp[OF measurable_component_singleton f] - by (auto intro!: borel_measurable_setprod simp: comp_def) + have borel: "?f \ borel_measurable (Pi\<^isub>M I M)" + using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto 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" + show "(\x. abs (?f x)) \ borel_measurable (Pi\<^isub>M I M)" using borel by auto - have "(\\<^isup>+x. ereal (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. ereal (abs (f i (x i)))) \P)" + have "(\\<^isup>+x. ereal (abs (?f x)) \Pi\<^isub>M I M) = (\\<^isup>+x. (\i\I. ereal (abs (f i (x i)))) \Pi\<^isub>M I M)" by (simp add: setprod_ereal abs_setprod) also have "\ = (\i\I. (\\<^isup>+x. ereal (abs (f i x)) \M i))" using f by (subst product_positive_integral_setprod) auto also have "\ < \" - using integrable[THEN M.integrable_abs] - by (simp add: setprod_PInf integrable_def M.positive_integral_positive) - finally show "(\\<^isup>+x. ereal (abs (?f x)) \P) \ \" by auto - have "(\\<^isup>+x. ereal (- abs (?f x)) \P) = (\\<^isup>+x. 0 \P)" + using integrable[THEN integrable_abs] + by (simp add: setprod_PInf integrable_def positive_integral_positive) + finally show "(\\<^isup>+x. ereal (abs (?f x)) \(Pi\<^isub>M I M)) \ \" by auto + have "(\\<^isup>+x. ereal (- abs (?f x)) \(Pi\<^isub>M I M)) = (\\<^isup>+x. 0 \(Pi\<^isub>M I M))" by (intro positive_integral_cong_pos) auto - then show "(\\<^isup>+x. ereal (- abs (?f x)) \P) \ \" by simp + then show "(\\<^isup>+x. ereal (- abs (?f x)) \(Pi\<^isub>M I M)) \ \" by simp qed ultimately show ?thesis by (rule integrable_abs_iff[THEN iffD1]) @@ -992,7 +999,69 @@ using `i \ I` by (auto intro!: setprod_cong) show ?case unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] - by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI) + by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI) qed -end \ No newline at end of file +lemma sigma_prod_algebra_sigma_eq: + fixes E :: "'i \ 'a set set" + assumes "finite I" + assumes S_mono: "\i. i \ I \ incseq (S i)" + and S_union: "\i. i \ I \ (\j. S i j) = space (M i)" + and S_in_E: "\i. i \ I \ range (S i) \ E i" + assumes E_closed: "\i. i \ I \ E i \ Pow (space (M i))" + and E_generates: "\i. i \ I \ sets (M i) = sigma_sets (space (M i)) (E i)" + defines "P == { Pi\<^isub>E I F | F. \i\I. F i \ E i }" + shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P" +proof + let ?P = "sigma (space (Pi\<^isub>M I M)) P" + have P_closed: "P \ Pow (space (Pi\<^isub>M I M))" + using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq) + then have space_P: "space ?P = (\\<^isub>E i\I. space (M i))" + by (simp add: space_PiM) + have "sets (PiM I M) = + sigma_sets (space ?P) {{f \ \\<^isub>E i\I. space (M i). f i \ A} |i A. i \ I \ A \ sets (M i)}" + using sets_PiM_single[of I M] by (simp add: space_P) + also have "\ \ sets (sigma (space (PiM I M)) P)" + proof (safe intro!: sigma_sets_subset) + fix i A assume "i \ I" and A: "A \ sets (M i)" + have "(\x. x i) \ measurable ?P (sigma (space (M i)) (E i))" + proof (subst measurable_iff_measure_of) + show "E i \ Pow (space (M i))" using `i \ I` by fact + from space_P `i \ I` show "(\x. x i) \ space ?P \ space (M i)" + by (auto simp: Pi_iff) + show "\A\E i. (\x. x i) -` A \ space ?P \ sets ?P" + proof + fix A assume A: "A \ E i" + then have "(\x. x i) -` A \ space ?P = (\\<^isub>E j\I. if i = j then A else space (M j))" + using E_closed `i \ I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm) + also have "\ = (\\<^isub>E j\I. \n. if i = j then A else S j n)" + by (intro PiE_cong) (simp add: S_union) + also have "\ = (\n. \\<^isub>E j\I. if i = j then A else S j n)" + using S_mono + by (subst Pi_UN[symmetric, OF `finite I`]) (auto simp: incseq_def) + also have "\ \ sets ?P" + proof (safe intro!: countable_UN) + fix n show "(\\<^isub>E j\I. if i = j then A else S j n) \ sets ?P" + using A S_in_E + by (simp add: P_closed) + (auto simp: P_def subset_eq intro!: exI[of _ "\j. if i = j then A else S j n"]) + qed + finally show "(\x. x i) -` A \ space ?P \ sets ?P" + using P_closed by simp + qed + qed + from measurable_sets[OF this, of A] A `i \ I` E_closed + have "(\x. x i) -` A \ space ?P \ sets ?P" + by (simp add: E_generates) + also have "(\x. x i) -` A \ space ?P = {f \ \\<^isub>E i\I. space (M i). f i \ A}" + using P_closed by (auto simp: space_PiM) + finally show "\ \ sets ?P" . + qed + finally show "sets (PiM I M) \ sigma_sets (space (PiM I M)) P" + by (simp add: P_closed) + show "sigma_sets (space (PiM I M)) P \ sets (PiM I M)" + using `finite I` + by (auto intro!: sigma_sets_subset simp: E_generates P_def) +qed + +end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Independent_Family.thy Mon Apr 23 12:14:35 2012 +0200 @@ -5,7 +5,7 @@ header {* Independent families of events, event sets, and random variables *} theory Independent_Family - imports Probability_Measure + imports Probability_Measure Infinite_Product_Measure begin lemma INT_decseq_offset: @@ -44,7 +44,7 @@ definition (in prob_space) "indep_var Ma A Mb B \ indep_vars (bool_case Ma Mb) (bool_case A B) UNIV" -lemma (in prob_space) indep_sets_cong[cong]: +lemma (in prob_space) indep_sets_cong: "I = J \ (\i. i \ I \ F i = G i) \ indep_sets F I \ indep_sets G J" by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+ @@ -135,7 +135,7 @@ lemma (in prob_space) indep_sets_dynkin: assumes indep: "indep_sets F I" - shows "indep_sets (\i. sets (dynkin \ space = space M, sets = F i \)) I" + shows "indep_sets (\i. dynkin (space M) (F i)) I" (is "indep_sets ?F I") proof (subst indep_sets_finite_index_sets, intro allI impI ballI) fix J assume "finite J" "J \ I" "J \ {}" @@ -193,7 +193,7 @@ qed qed } note indep_sets_insert = this - have "dynkin_system \ space = space M, sets = ?D \" + have "dynkin_system (space M) ?D" proof (rule dynkin_systemI', simp_all cong del: indep_sets_cong, safe) show "indep_sets (G(j := {{}})) K" by (rule indep_sets_insert) auto @@ -206,7 +206,7 @@ using G by auto have "prob ((\j\J. A j) \ (space M - X)) = prob ((\j\J. A j) - (\i\insert j J. (A(j := X)) i))" - using A_sets sets_into_space X `J \ {}` + using A_sets sets_into_space[of _ M] X `J \ {}` by (auto intro!: arg_cong[where f=prob] split: split_if_asm) also have "\ = prob (\j\J. A j) - prob (\i\insert j J. (A(j := X)) i)" using J `J \ {}` `j \ J` A_sets X sets_into_space @@ -264,9 +264,8 @@ by (auto dest!: sums_unique) qed (insert F, auto) qed (insert sets_into_space, auto) - then have mono: "sets (dynkin \space = space M, sets = G j\) \ - sets \space = space M, sets = {E \ events. indep_sets (G(j := {E})) K}\" - proof (rule dynkin_system.dynkin_subset, simp_all cong del: indep_sets_cong, safe) + then have mono: "dynkin (space M) (G j) \ {E \ events. indep_sets (G(j := {E})) K}" + proof (rule dynkin_system.dynkin_subset, safe) fix X assume "X \ G j" then show "X \ events" using G `j \ K` by auto from `indep_sets G K` @@ -292,20 +291,20 @@ by (intro indep_setsD[OF G(1)]) auto qed qed - then have "indep_sets (G(j:=sets (dynkin \space = space M, sets = G j\))) K" + then have "indep_sets (G(j := dynkin (space M) (G j))) K" by (rule indep_sets_mono_sets) (insert mono, auto) then show ?case by (rule indep_sets_mono_sets) (insert `j \ K` `j \ J`, auto simp: G_def) qed (insert `indep_sets F K`, simp) } from this[OF `indep_sets F J` `finite J` subset_refl] - show "indep_sets (\i. sets (dynkin \ space = space M, sets = F i \)) J" + show "indep_sets ?F J" by (rule indep_sets_mono_sets) auto qed lemma (in prob_space) indep_sets_sigma: assumes indep: "indep_sets F I" - assumes stable: "\i. i \ I \ Int_stable \ space = space M, sets = F i \" - shows "indep_sets (\i. sets (sigma \ space = space M, sets = F i \)) I" + assumes stable: "\i. i \ I \ Int_stable (F i)" + shows "indep_sets (\i. sigma_sets (space M) (F i)) I" proof - from indep_sets_dynkin[OF indep] show ?thesis @@ -316,18 +315,12 @@ qed qed -lemma (in prob_space) indep_sets_sigma_sets: - assumes "indep_sets F I" - assumes "\i. i \ I \ Int_stable \ space = space M, sets = F i \" - shows "indep_sets (\i. sigma_sets (space M) (F i)) I" - using indep_sets_sigma[OF assms] by (simp add: sets_sigma) - lemma (in prob_space) indep_sets_sigma_sets_iff: - assumes "\i. i \ I \ Int_stable \ space = space M, sets = F i \" + assumes "\i. i \ I \ Int_stable (F i)" shows "indep_sets (\i. sigma_sets (space M) (F i)) I \ indep_sets F I" proof assume "indep_sets F I" then show "indep_sets (\i. sigma_sets (space M) (F i)) I" - by (rule indep_sets_sigma_sets) fact + by (rule indep_sets_sigma) fact next assume "indep_sets (\i. sigma_sets (space M) (F i)) I" then show "indep_sets F I" by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic) @@ -361,15 +354,14 @@ lemma (in prob_space) indep_set_sigma_sets: assumes "indep_set A B" - assumes A: "Int_stable \ space = space M, sets = A \" - assumes B: "Int_stable \ space = space M, sets = B \" + assumes A: "Int_stable A" and B: "Int_stable B" shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)" proof - have "indep_sets (\i. sigma_sets (space M) (case i of True \ A | False \ B)) UNIV" - proof (rule indep_sets_sigma_sets) + proof (rule indep_sets_sigma) show "indep_sets (bool_case A B) UNIV" by (rule `indep_set A B`[unfolded indep_set_def]) - fix i show "Int_stable \space = space M, sets = case i of True \ A | False \ B\" + fix i show "Int_stable (case i of True \ A | False \ B)" using A B by (cases i) auto qed then show ?thesis @@ -380,7 +372,7 @@ lemma (in prob_space) indep_sets_collect_sigma: fixes I :: "'j \ 'i set" and J :: "'j set" and E :: "'i \ 'a set set" assumes indep: "indep_sets E (\j\J. I j)" - assumes Int_stable: "\i j. j \ J \ i \ I j \ Int_stable \space = space M, sets = E i\" + assumes Int_stable: "\i j. j \ J \ i \ I j \ Int_stable (E i)" assumes disjoint: "disjoint_family_on I J" shows "indep_sets (\j. sigma_sets (space M) (\i\I j. E i)) J" proof - @@ -389,30 +381,29 @@ from indep have E: "\j i. j \ J \ i \ I j \ E i \ events" unfolding indep_sets_def by auto { fix j - let ?S = "sigma \ space = space M, sets = (\i\I j. E i) \" + let ?S = "sigma_sets (space M) (\i\I j. E i)" assume "j \ J" - from E[OF this] interpret S: sigma_algebra ?S - using sets_into_space by (intro sigma_algebra_sigma) auto + from E[OF this] interpret S: sigma_algebra "space M" ?S + using sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto have "sigma_sets (space M) (\i\I j. E i) = sigma_sets (space M) (?E j)" proof (rule sigma_sets_eqI) fix A assume "A \ (\i\I j. E i)" then guess i .. then show "A \ sigma_sets (space M) (?E j)" - by (auto intro!: sigma_sets.intros exI[of _ "{i}"] exI[of _ "\i. A"]) + by (auto intro!: sigma_sets.intros(2-) exI[of _ "{i}"] exI[of _ "\i. A"]) next fix A assume "A \ ?E j" then obtain E' K where "finite K" "K \ {}" "K \ I j" "\k. k \ K \ E' k \ E k" and A: "A = (\k\K. E' k)" by auto - then have "A \ sets ?S" unfolding A - by (safe intro!: S.finite_INT) - (auto simp: sets_sigma intro!: sigma_sets.Basic) + then have "A \ ?S" unfolding A + by (safe intro!: S.finite_INT) auto then show "A \ sigma_sets (space M) (\i\I j. E i)" - by (simp add: sets_sigma) + by simp qed } moreover have "indep_sets (\j. sigma_sets (space M) (?E j)) J" - proof (rule indep_sets_sigma_sets) + proof (rule indep_sets_sigma) show "indep_sets ?E J" proof (intro indep_setsI) fix j assume "j \ J" with E show "?E j \ events" by (force intro!: finite_INT) @@ -460,7 +451,7 @@ qed next fix j assume "j \ J" - show "Int_stable \ space = space M, sets = ?E j \" + show "Int_stable (?E j)" proof (rule Int_stableI) fix a assume "a \ ?E j" then obtain Ka Ea where a: "a = (\k\Ka. Ea k)" "finite Ka" "Ka \ {}" "Ka \ I j" "\k. k\Ka \ Ea k \ E k" by auto @@ -482,12 +473,12 @@ lemma (in prob_space) terminal_events_sets: assumes A: "\i. A i \ events" - assumes "\i::nat. sigma_algebra \space = space M, sets = A i\" + assumes "\i::nat. sigma_algebra (space M) (A i)" assumes X: "X \ terminal_events A" shows "X \ events" proof - let ?A = "(\n. sigma_sets (space M) (UNION {n..} A))" - interpret A: sigma_algebra "\space = space M, sets = A i\" for i by fact + interpret A: sigma_algebra "space M" "A i" for i by fact from X have "\n. X \ sigma_sets (space M) (UNION {n..} A)" by (auto simp: terminal_events_def) from this[of 0] have "X \ sigma_sets (space M) (UNION UNIV A)" by simp then show "X \ events" @@ -495,12 +486,12 @@ qed lemma (in prob_space) sigma_algebra_terminal_events: - assumes "\i::nat. sigma_algebra \space = space M, sets = A i\" - shows "sigma_algebra \ space = space M, sets = terminal_events A \" + assumes "\i::nat. sigma_algebra (space M) (A i)" + shows "sigma_algebra (space M) (terminal_events A)" unfolding terminal_events_def proof (simp add: sigma_algebra_iff2, safe) let ?A = "(\n. sigma_sets (space M) (UNION {n..} A))" - interpret A: sigma_algebra "\space = space M, sets = A i\" for i by fact + interpret A: sigma_algebra "space M" "A i" for i by fact { fix X x assume "X \ ?A" "x \ X" then have "\n. X \ sigma_sets (space M) (UNION {n..} A)" by auto from this[of 0] have "X \ sigma_sets (space M) (UNION UNIV A)" by simp @@ -515,29 +506,29 @@ lemma (in prob_space) kolmogorov_0_1_law: fixes A :: "nat \ 'a set set" assumes A: "\i. A i \ events" - assumes "\i::nat. sigma_algebra \space = space M, sets = A i\" + assumes "\i::nat. sigma_algebra (space M) (A i)" assumes indep: "indep_sets A UNIV" and X: "X \ terminal_events A" shows "prob X = 0 \ prob X = 1" proof - - let ?D = "\ space = space M, sets = {D \ events. prob (X \ D) = prob X * prob D} \" - interpret A: sigma_algebra "\space = space M, sets = A i\" for i by fact - interpret T: sigma_algebra "\ space = space M, sets = terminal_events A \" + let ?D = "{D \ events. prob (X \ D) = prob X * prob D}" + interpret A: sigma_algebra "space M" "A i" for i by fact + interpret T: sigma_algebra "space M" "terminal_events A" by (rule sigma_algebra_terminal_events) fact have "X \ space M" using T.space_closed X by auto have X_in: "X \ events" by (rule terminal_events_sets) fact+ - interpret D: dynkin_system ?D + interpret D: dynkin_system "space M" ?D proof (rule dynkin_systemI) - fix D assume "D \ sets ?D" then show "D \ space ?D" + fix D assume "D \ ?D" then show "D \ space M" using sets_into_space by auto next - show "space ?D \ sets ?D" + show "space M \ ?D" using prob_space `X \ space M` by (simp add: Int_absorb2) next - fix A assume A: "A \ sets ?D" + fix A assume A: "A \ ?D" have "prob (X \ (space M - A)) = prob (X - (X \ A))" using `X \ space M` by (auto intro!: arg_cong[where f=prob]) also have "\ = prob X - prob (X \ A)" @@ -547,10 +538,10 @@ also have "\ = prob X * prob (space M - A)" using X_in A sets_into_space by (subst finite_measure_Diff) (auto simp: field_simps) - finally show "space ?D - A \ sets ?D" + finally show "space M - A \ ?D" using A `X \ space M` by auto next - fix F :: "nat \ 'a set" assume dis: "disjoint_family F" and "range F \ sets ?D" + fix F :: "nat \ 'a set" assume dis: "disjoint_family F" and "range F \ ?D" then have F: "range F \ events" "\i. prob (X \ F i) = prob X * prob (F i)" by auto have "(\i. prob (X \ F i)) sums prob (\i. X \ F i)" @@ -566,7 +557,7 @@ by (intro sums_mult finite_measure_UNION F dis) ultimately have "prob (X \ (\i. F i)) = prob X * prob (\i. F i)" by (auto dest!: sums_unique) - with F show "(\i. F i) \ sets ?D" + with F show "(\i. F i) \ ?D" by auto qed @@ -579,7 +570,7 @@ show "disjoint_family (bool_case {..n} {Suc n..})" unfolding disjoint_family_on_def by (auto split: bool.split) fix m - show "Int_stable \space = space M, sets = A m\" + show "Int_stable (A m)" unfolding Int_stable_def using A.Int by auto qed also have "(\b. sigma_sets (space M) (\m\bool_case {..n} {Suc n..} b. A m)) = @@ -588,7 +579,7 @@ finally have indep: "indep_set (sigma_sets (space M) (\m\{..n}. A m)) (sigma_sets (space M) (\m\{Suc n..}. A m))" unfolding indep_set_def by simp - have "sigma_sets (space M) (\m\{..n}. A m) \ sets ?D" + have "sigma_sets (space M) (\m\{..n}. A m) \ ?D" proof (simp add: subset_eq, rule) fix D assume D: "D \ sigma_sets (space M) (\m\{..n}. A m)" have "X \ sigma_sets (space M) (\m\{Suc n..}. A m)" @@ -597,22 +588,27 @@ show "D \ events \ prob (X \ D) = prob X * prob D" by (auto simp add: ac_simps) qed } - then have "(\n. sigma_sets (space M) (\m\{..n}. A m)) \ sets ?D" (is "?A \ _") + then have "(\n. sigma_sets (space M) (\m\{..n}. A m)) \ ?D" (is "?A \ _") by auto - have "sigma \ space = space M, sets = ?A \ = - dynkin \ space = space M, sets = ?A \" (is "sigma ?UA = dynkin ?UA") + note `X \ terminal_events A` + also { + have "\n. sigma_sets (space M) (\i\{n..}. A i) \ sigma_sets (space M) ?A" + by (intro sigma_sets_subseteq UN_mono) auto + then have "terminal_events A \ sigma_sets (space M) ?A" + unfolding terminal_events_def by auto } + also have "sigma_sets (space M) ?A = dynkin (space M) ?A" proof (rule sigma_eq_dynkin) { fix B n assume "B \ sigma_sets (space M) (\m\{..n}. A m)" then have "B \ space M" - by induct (insert A sets_into_space, auto) } - then show "sets ?UA \ Pow (space ?UA)" by auto - show "Int_stable ?UA" + by induct (insert A sets_into_space[of _ M], auto) } + then show "?A \ Pow (space M)" by auto + show "Int_stable ?A" proof (rule Int_stableI) fix a assume "a \ ?A" then guess n .. note a = this fix b assume "b \ ?A" then guess m .. note b = this - interpret Amn: sigma_algebra "sigma \space = space M, sets = (\i\{..max m n}. A i)\" - using A sets_into_space by (intro sigma_algebra_sigma) auto + interpret Amn: sigma_algebra "space M" "sigma_sets (space M) (\i\{..max m n}. A i)" + using A sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto have "sigma_sets (space M) (\i\{..n}. A i) \ sigma_sets (space M) (\i\{..max m n}. A i)" by (intro sigma_sets_subseteq UN_mono) auto with a have "a \ sigma_sets (space M) (\i\{..max m n}. A i)" by auto @@ -621,23 +617,13 @@ by (intro sigma_sets_subseteq UN_mono) auto with b have "b \ sigma_sets (space M) (\i\{..max m n}. A i)" by auto ultimately have "a \ b \ sigma_sets (space M) (\i\{..max m n}. A i)" - using Amn.Int[of a b] by (simp add: sets_sigma) + using Amn.Int[of a b] by simp then show "a \ b \ (\n. sigma_sets (space M) (\i\{..n}. A i))" by auto qed qed - moreover have "sets (dynkin ?UA) \ sets ?D" - proof (rule D.dynkin_subset) - show "sets ?UA \ sets ?D" using `?A \ sets ?D` by auto - qed simp - ultimately have "sets (sigma ?UA) \ sets ?D" by simp - moreover - have "\n. sigma_sets (space M) (\i\{n..}. A i) \ sigma_sets (space M) ?A" - by (intro sigma_sets_subseteq UN_mono) (auto intro: sigma_sets.Basic) - then have "terminal_events A \ sets (sigma ?UA)" - unfolding sets_sigma terminal_events_def by auto - moreover note `X \ terminal_events A` - ultimately have "X \ sets ?D" by auto - then show ?thesis by auto + also have "dynkin (space M) ?A \ ?D" + using `?A \ ?D` by (auto intro!: D.dynkin_subset) + finally show ?thesis by auto qed lemma (in prob_space) borel_0_1_law: @@ -648,14 +634,14 @@ show "\i. sigma_sets (space M) {F i} \ events" using F(1) sets_into_space by (subst sigma_sets_singleton) auto - { fix i show "sigma_algebra \space = space M, sets = sigma_sets (space M) {F i}\" - using sigma_algebra_sigma[of "\space = space M, sets = {F i}\"] F sets_into_space - by (auto simp add: sigma_def) } + { fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})" + using sigma_algebra_sigma_sets[of "{F i}" "space M"] F sets_into_space + by auto } show "indep_sets (\i. sigma_sets (space M) {F i}) UNIV" - proof (rule indep_sets_sigma_sets) + proof (rule indep_sets_sigma) show "indep_sets (\i. {F i}) UNIV" unfolding indep_sets_singleton_iff_indep_events by fact - fix i show "Int_stable \space = space M, sets = {F i}\" + fix i show "Int_stable {F i}" unfolding Int_stable_def by simp qed let ?Q = "\n. \i\{n..}. F i" @@ -663,17 +649,17 @@ unfolding terminal_events_def proof fix j - interpret S: sigma_algebra "sigma \ space = space M, sets = (\i\{j..}. sigma_sets (space M) {F i})\" + interpret S: sigma_algebra "space M" "sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" using order_trans[OF F(1) space_closed] - by (intro sigma_algebra_sigma) (simp add: sigma_sets_singleton subset_eq) + by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq) have "(\n. ?Q n) = (\n\{j..}. ?Q n)" by (intro decseq_SucI INT_decseq_offset UN_mono) auto - also have "\ \ sets (sigma \ space = space M, sets = (\i\{j..}. sigma_sets (space M) {F i})\)" + also have "\ \ sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" using order_trans[OF F(1) space_closed] by (safe intro!: S.countable_INT S.countable_UN) - (auto simp: sets_sigma sigma_sets_singleton intro!: sigma_sets.Basic bexI) + (auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI) finally show "(\n. ?Q n) \ sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" - by (simp add: sets_sigma) + by simp qed qed @@ -710,84 +696,84 @@ lemma (in prob_space) indep_vars_finite: fixes I :: "'i set" assumes I: "I \ {}" "finite I" - and rv: "\i. i \ I \ random_variable (sigma (M' i)) (X i)" - and Int_stable: "\i. i \ I \ Int_stable (M' i)" - and space: "\i. i \ I \ space (M' i) \ sets (M' i)" - shows "indep_vars (\i. sigma (M' i)) X I \ - (\A\(\ i\I. sets (M' i)). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M)))" + and M': "\i. i \ I \ sets (M' i) = sigma_sets (space (M' i)) (E i)" + and rv: "\i. i \ I \ random_variable (M' i) (X i)" + and Int_stable: "\i. i \ I \ Int_stable (E i)" + and space: "\i. i \ I \ space (M' i) \ E i" and closed: "\i. i \ I \ E i \ Pow (space (M' i))" + shows "indep_vars M' X I \ + (\A\(\ i\I. E i). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M)))" proof - from rv have X: "\i. i \ I \ X i \ space M \ space (M' i)" unfolding measurable_def by simp { fix i assume "i\I" - have "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (sigma (M' i))} - = sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" - unfolding sigma_sets_vimage_commute[OF X, OF `i \ I`] + from closed[OF `i \ I`] + have "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)} + = sigma_sets (space M) {X i -` A \ space M |A. A \ E i}" + unfolding sigma_sets_vimage_commute[OF X, OF `i \ I`, symmetric] M'[OF `i \ I`] by (subst sigma_sets_sigma_sets_eq) auto } - note this[simp] + note sigma_sets_X = this { fix i assume "i\I" - have "Int_stable \space = space M, sets = {X i -` A \ space M |A. A \ sets (M' i)}\" + have "Int_stable {X i -` A \ space M |A. A \ E i}" proof (rule Int_stableI) - fix a assume "a \ {X i -` A \ space M |A. A \ sets (M' i)}" - then obtain A where "a = X i -` A \ space M" "A \ sets (M' i)" by auto + fix a assume "a \ {X i -` A \ space M |A. A \ E i}" + then obtain A where "a = X i -` A \ space M" "A \ E i" by auto moreover - fix b assume "b \ {X i -` A \ space M |A. A \ sets (M' i)}" - then obtain B where "b = X i -` B \ space M" "B \ sets (M' i)" by auto + fix b assume "b \ {X i -` A \ space M |A. A \ E i}" + then obtain B where "b = X i -` B \ space M" "B \ E i" by auto moreover have "(X i -` A \ space M) \ (X i -` B \ space M) = X i -` (A \ B) \ space M" by auto moreover note Int_stable[OF `i \ I`] ultimately - show "a \ b \ {X i -` A \ space M |A. A \ sets (M' i)}" + show "a \ b \ {X i -` A \ space M |A. A \ E i}" by (auto simp del: vimage_Int intro!: exI[of _ "A \ B"] dest: Int_stableD) qed } - note indep_sets_sigma_sets_iff[OF this, simp] + note indep_sets_X = indep_sets_sigma_sets_iff[OF this] { fix i assume "i \ I" - { fix A assume "A \ sets (M' i)" - then have "A \ sets (sigma (M' i))" by (auto simp: sets_sigma intro: sigma_sets.Basic) + { fix A assume "A \ E i" + with M'[OF `i \ I`] have "A \ sets (M' i)" by auto moreover - from rv[OF `i\I`] have "X i \ measurable M (sigma (M' i))" by auto + from rv[OF `i\I`] have "X i \ measurable M (M' i)" by auto ultimately have "X i -` A \ space M \ sets M" by (auto intro: measurable_sets) } with X[OF `i\I`] space[OF `i\I`] - have "{X i -` A \ space M |A. A \ sets (M' i)} \ events" - "space M \ {X i -` A \ space M |A. A \ sets (M' i)}" + have "{X i -` A \ space M |A. A \ E i} \ events" + "space M \ {X i -` A \ space M |A. A \ E i}" by (auto intro!: exI[of _ "space (M' i)"]) } - note indep_sets_finite[OF I this, simp] + note indep_sets_finite_X = indep_sets_finite[OF I this] - have "(\A\\ i\I. {X i -` A \ space M |A. A \ sets (M' i)}. prob (INTER I A) = (\j\I. prob (A j))) = - (\A\\ i\I. sets (M' i). prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M)))" + have "(\A\\ i\I. {X i -` A \ space M |A. A \ E i}. prob (INTER I A) = (\j\I. prob (A j))) = + (\A\\ i\I. E i. prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M)))" (is "?L = ?R") proof safe - fix A assume ?L and A: "A \ (\ i\I. sets (M' i))" + fix A assume ?L and A: "A \ (\ i\I. E i)" from `?L`[THEN bspec, of "\i. X i -` A i \ space M"] A `I \ {}` show "prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M))" by (auto simp add: Pi_iff) next - fix A assume ?R and A: "A \ (\ i\I. {X i -` A \ space M |A. A \ sets (M' i)})" - from A have "\i\I. \B. A i = X i -` B \ space M \ B \ sets (M' i)" by auto + fix A assume ?R and A: "A \ (\ i\I. {X i -` A \ space M |A. A \ E i})" + from A have "\i\I. \B. A i = X i -` B \ space M \ B \ E i" by auto from bchoice[OF this] obtain B where B: "\i\I. A i = X i -` B i \ space M" - "B \ (\ i\I. sets (M' i))" by auto + "B \ (\ i\I. E i)" by auto from `?R`[THEN bspec, OF B(2)] B(1) `I \ {}` show "prob (INTER I A) = (\j\I. prob (A j))" by simp qed then show ?thesis using `I \ {}` - by (simp add: rv indep_vars_def) + by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong) qed lemma (in prob_space) indep_vars_compose: assumes "indep_vars M' X I" - assumes rv: - "\i. i \ I \ sigma_algebra (N i)" - "\i. i \ I \ Y i \ measurable (M' i) (N i)" + assumes rv: "\i. i \ I \ Y i \ measurable (M' i) (N i)" shows "indep_vars N (\i. Y i \ X i) I" unfolding indep_vars_def proof from rv `indep_vars M' X I` show "\i\I. random_variable (N i) (Y i \ X i)" - by (auto intro!: measurable_comp simp: indep_vars_def) + by (auto simp: indep_vars_def) have "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" using `indep_vars M' X I` by (simp add: indep_vars_def) @@ -806,7 +792,7 @@ qed qed -lemma (in prob_space) indep_varsD: +lemma (in prob_space) indep_varsD_finite: assumes X: "indep_vars M' X I" assumes I: "I \ {}" "finite I" "\i. i \ I \ A i \ sets (M' i)" shows "prob (\i\I. X i -` A i \ space M) = (\i\I. prob (X i -` A i \ space M))" @@ -815,96 +801,134 @@ using X by (auto simp: indep_vars_def) show "I \ I" "I \ {}" "finite I" using I by auto show "\i\I. X i -` A i \ space M \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" - using I by (auto intro: sigma_sets.Basic) + using I by auto qed -lemma (in prob_space) indep_distribution_eq_measure: - assumes I: "I \ {}" "finite I" +lemma (in prob_space) indep_varsD: + assumes X: "indep_vars M' X I" + assumes I: "J \ {}" "finite J" "J \ I" "\i. i \ J \ A i \ sets (M' i)" + shows "prob (\i\J. X i -` A i \ space M) = (\i\J. prob (X i -` A i \ space M))" +proof (rule indep_setsD) + show "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" + using X by (auto simp: indep_vars_def) + show "\i\J. X i -` A i \ space M \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" + using I by auto +qed fact+ + +lemma prod_algebra_cong: + assumes "I = J" and sets: "(\i. i \ I \ sets (M i) = sets (N i))" + shows "prod_algebra I M = prod_algebra J N" +proof - + have space: "\i. i \ I \ space (M i) = space (N i)" + using sets_eq_imp_space_eq[OF sets] by auto + with sets show ?thesis unfolding `I = J` + by (intro antisym prod_algebra_mono) auto +qed + +lemma space_in_prod_algebra: + "(\\<^isub>E i\I. space (M i)) \ prod_algebra I M" +proof cases + assume "I = {}" then show ?thesis + by (auto simp add: prod_algebra_def image_iff prod_emb_def) +next + assume "I \ {}" + then obtain i where "i \ I" by auto + then have "(\\<^isub>E i\I. space (M i)) = prod_emb I M {i} (\\<^isub>E i\{i}. space (M i))" + by (auto simp: prod_emb_def Pi_iff) + also have "\ \ prod_algebra I M" + using `i \ I` by (intro prod_algebraI) auto + finally show ?thesis . +qed + +lemma (in prob_space) indep_vars_iff_distr_eq_PiM: + fixes I :: "'i set" and X :: "'i \ 'a \ 'b" + assumes "I \ {}" assumes rv: "\i. random_variable (M' i) (X i)" shows "indep_vars M' X I \ - (\A\sets (\\<^isub>M i\I. (M' i \ measure := ereal\distribution (X i) \)). - distribution (\x. \i\I. X i x) A = - finite_measure.\' (\\<^isub>M i\I. (M' i \ measure := ereal\distribution (X i) \)) A)" - (is "_ \ (\X\_. distribution ?D X = finite_measure.\' (Pi\<^isub>M I ?M) X)") + distr M (\\<^isub>M i\I. M' i) (\x. \i\I. X i x) = (\\<^isub>M i\I. distr M (M' i) (X i))" proof - - interpret M': prob_space "?M i" for i - using rv by (rule distribution_prob_space) - interpret P: finite_product_prob_space ?M I - proof qed fact + let ?P = "\\<^isub>M i\I. M' i" + let ?X = "\x. \i\I. X i x" + let ?D = "distr M ?P ?X" + have X: "random_variable ?P ?X" by (intro measurable_restrict rv) + interpret D: prob_space ?D by (intro prob_space_distr X) - let ?D' = "(Pi\<^isub>M I ?M) \ measure := ereal \ distribution ?D \" - have "random_variable P.P ?D" - using `finite I` rv by (intro random_variable_restrict) auto - then interpret D: prob_space ?D' - by (rule distribution_prob_space) - + let ?D' = "\i. distr M (M' i) (X i)" + let ?P' = "\\<^isub>M i\I. distr M (M' i) (X i)" + interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv) + interpret P: product_prob_space ?D' I .. + show ?thesis - proof (intro iffI ballI) + proof assume "indep_vars M' X I" - fix A assume "A \ sets P.P" - moreover - have "D.prob A = P.prob A" - proof (rule prob_space_unique_Int_stable) - show "prob_space ?D'" by unfold_locales - show "prob_space (Pi\<^isub>M I ?M)" by unfold_locales - show "Int_stable P.G" using M'.Int - by (intro Int_stable_product_algebra_generator) (simp add: Int_stable_def) - show "space P.G \ sets P.G" - using M'.top by (simp add: product_algebra_generator_def) - show "space ?D' = space P.G" "sets ?D' = sets (sigma P.G)" - by (simp_all add: product_algebra_def product_algebra_generator_def sets_sigma) - show "space P.P = space P.G" "sets P.P = sets (sigma P.G)" - by (simp_all add: product_algebra_def) - show "A \ sets (sigma P.G)" - using `A \ sets P.P` by (simp add: product_algebra_def) + show "?D = ?P'" + proof (rule measure_eqI_generator_eq) + show "Int_stable (prod_algebra I M')" + by (rule Int_stable_prod_algebra) + show "prod_algebra I M' \ Pow (space ?P)" + using prod_algebra_sets_into_space by (simp add: space_PiM) + show "sets ?D = sigma_sets (space ?P) (prod_algebra I M')" + by (simp add: sets_PiM space_PiM) + show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')" + by (simp add: sets_PiM space_PiM cong: prod_algebra_cong) + let ?A = "\i. \\<^isub>E i\I. space (M' i)" + show "range ?A \ prod_algebra I M'" "incseq ?A" "(\i. ?A i) = space (Pi\<^isub>M I M')" + by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong) + { fix i show "emeasure ?D (\\<^isub>E i\I. space (M' i)) \ \" by auto } + next + fix E assume E: "E \ prod_algebra I M'" + from prod_algebraE[OF E] guess J Y . note J = this - fix E assume E: "E \ sets P.G" - then have "E \ sets P.P" - by (simp add: sets_sigma sigma_sets.Basic product_algebra_def) - then have "D.prob E = distribution ?D E" - unfolding D.\'_def by simp - also - from E obtain F where "E = Pi\<^isub>E I F" and F: "\i. i \ I \ F i \ sets (M' i)" - by (auto simp: product_algebra_generator_def) - with `I \ {}` have "distribution ?D E = prob (\i\I. X i -` F i \ space M)" - using `I \ {}` by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def) - also have "\ = (\i\I. prob (X i -` F i \ space M))" - using `indep_vars M' X I` I F by (rule indep_varsD) - also have "\ = P.prob E" - using F by (simp add: `E = Pi\<^isub>E I F` P.prob_times M'.\'_def distribution_def) - finally show "D.prob E = P.prob E" . + from E have "E \ sets ?P" by (auto simp: sets_PiM) + then have "emeasure ?D E = emeasure M (?X -` E \ space M)" + by (simp add: emeasure_distr X) + also have "?X -` E \ space M = (\i\J. X i -` Y i \ space M)" + using J `I \ {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm) + also have "emeasure M (\i\J. X i -` Y i \ space M) = (\ i\J. emeasure M (X i -` Y i \ space M))" + using `indep_vars M' X I` J `I \ {}` using indep_varsD[of M' X I J] + by (auto simp: emeasure_eq_measure setprod_ereal) + also have "\ = (\ i\J. emeasure (?D' i) (Y i))" + using rv J by (simp add: emeasure_distr) + also have "\ = emeasure ?P' E" + using P.emeasure_PiM_emb[of J Y] J by (simp add: prod_emb_def) + finally show "emeasure ?D E = emeasure ?P' E" . qed - ultimately show "distribution ?D A = P.prob A" - by (simp add: D.\'_def) next - assume eq: "\A\sets P.P. distribution ?D A = P.prob A" - have [simp]: "\i. sigma (M' i) = M' i" - using rv by (intro sigma_algebra.sigma_eq) simp - have "indep_vars (\i. sigma (M' i)) X I" - proof (subst indep_vars_finite[OF I]) - fix i assume [simp]: "i \ I" - show "random_variable (sigma (M' i)) (X i)" - using rv[of i] by simp - show "Int_stable (M' i)" "space (M' i) \ sets (M' i)" - using M'.Int[of _ i] M'.top by (auto simp: Int_stable_def) + assume "?D = ?P'" + show "indep_vars M' X I" unfolding indep_vars_def + proof (intro conjI indep_setsI ballI rv) + fix i show "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)} \ events" + by (auto intro!: sigma_sets_subset measurable_sets rv) next - show "\A\\ i\I. sets (M' i). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M))" + fix J Y' assume J: "J \ {}" "J \ I" "finite J" + assume Y': "\j\J. Y' j \ sigma_sets (space M) {X j -` A \ space M |A. A \ sets (M' j)}" + have "\j\J. \Y. Y' j = X j -` Y \ space M \ Y \ sets (M' j)" proof - fix A assume A: "A \ (\ i\I. sets (M' i))" - then have A_in_P: "(Pi\<^isub>E I A) \ sets P.P" - by (auto intro!: product_algebraI) - have "prob (\j\I. X j -` A j \ space M) = distribution ?D (Pi\<^isub>E I A)" - using `I \ {}`by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def) - also have "\ = P.prob (Pi\<^isub>E I A)" using A_in_P eq by simp - also have "\ = (\i\I. M'.prob i (A i))" - using A by (intro P.prob_times) auto - also have "\ = (\i\I. prob (X i -` A i \ space M))" - using A by (auto intro!: setprod_cong simp: M'.\'_def Pi_iff distribution_def) - finally show "prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M))" . + fix j assume "j \ J" + from Y'[rule_format, OF this] rv[of j] + show "\Y. Y' j = X j -` Y \ space M \ Y \ sets (M' j)" + by (subst (asm) sigma_sets_vimage_commute[symmetric, of _ _ "space (M' j)"]) + (auto dest: measurable_space simp: sigma_sets_eq) qed + from bchoice[OF this] obtain Y where + Y: "\j. j \ J \ Y' j = X j -` Y j \ space M" "\j. j \ J \ Y j \ sets (M' j)" by auto + let ?E = "prod_emb I M' J (Pi\<^isub>E J Y)" + from Y have "(\j\J. Y' j) = ?X -` ?E \ space M" + using J `I \ {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm) + then have "emeasure M (\j\J. Y' j) = emeasure M (?X -` ?E \ space M)" + by simp + also have "\ = emeasure ?D ?E" + using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto + also have "\ = emeasure ?P' ?E" + using `?D = ?P'` by simp + also have "\ = (\ i\J. emeasure (?D' i) (Y i))" + using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def) + also have "\ = (\ i\J. emeasure M (Y' i))" + using rv J Y by (simp add: emeasure_distr) + finally have "emeasure M (\j\J. Y' j) = (\ i\J. emeasure M (Y' i))" . + then show "prob (\j\J. Y' j) = (\ i\J. prob (Y' i))" + by (auto simp: emeasure_eq_measure setprod_ereal) qed - then show "indep_vars M' X I" - by simp qed qed @@ -936,56 +960,188 @@ unfolding UNIV_bool by auto qed -lemma (in prob_space) indep_var_distributionD: - assumes indep: "indep_var S X T Y" - defines "P \ S\measure := ereal\distribution X\ \\<^isub>M T\measure := ereal\distribution Y\" - assumes "A \ sets P" - shows "joint_distribution X Y A = finite_measure.\' P A" -proof - - from indep have rvs: "random_variable S X" "random_variable T Y" - by (blast dest: indep_var_rv1 indep_var_rv2)+ +lemma measurable_bool_case[simp, intro]: + "(\(x, y). bool_case x y) \ measurable (M \\<^isub>M N) (Pi\<^isub>M UNIV (bool_case M N))" + (is "?f \ measurable ?B ?P") +proof (rule measurable_PiM_single) + show "?f \ space ?B \ (\\<^isub>E i\UNIV. space (bool_case M N i))" + by (auto simp: space_pair_measure extensional_def split: bool.split) + fix i A assume "A \ sets (case i of True \ M | False \ N)" + moreover then have "{\ \ space (M \\<^isub>M N). prod_case bool_case \ i \ A} + = (case i of True \ A \ space N | False \ space M \ A)" + by (auto simp: space_pair_measure split: bool.split dest!: sets_into_space) + ultimately show "{\ \ space (M \\<^isub>M N). prod_case bool_case \ i \ A} \ sets ?B" + by (auto split: bool.split) +qed + +lemma borel_measurable_indicator': + "A \ sets N \ f \ measurable M N \ (\x. indicator A (f x)) \ borel_measurable M" + using measurable_comp[OF _ borel_measurable_indicator, of f M N A] by (auto simp add: comp_def) + +lemma (in product_sigma_finite) distr_component: + "distr (M i) (Pi\<^isub>M {i} M) (\x. \i\{i}. x) = Pi\<^isub>M {i} M" (is "?D = ?P") +proof (intro measure_eqI[symmetric]) + interpret I: finite_product_sigma_finite M "{i}" by default simp + + have eq: "\x. x \ extensional {i} \ (\j\{i}. x i) = x" + by (auto simp: extensional_def restrict_def) + + fix A assume A: "A \ sets ?P" + then have "emeasure ?P A = (\\<^isup>+x. indicator A x \?P)" + by simp + also have "\ = (\\<^isup>+x. indicator ((\x. \i\{i}. x) -` A \ space (M i)) x \M i)" + apply (subst product_positive_integral_singleton[symmetric]) + apply (force intro!: measurable_restrict measurable_sets A) + apply (auto intro!: positive_integral_cong simp: space_PiM indicator_def simp: eq) + done + also have "\ = emeasure (M i) ((\x. \i\{i}. x) -` A \ space (M i))" + by (force intro!: measurable_restrict measurable_sets A positive_integral_indicator) + also have "\ = emeasure ?D A" + using A by (auto intro!: emeasure_distr[symmetric] measurable_restrict) + finally show "emeasure (Pi\<^isub>M {i} M) A = emeasure ?D A" . +qed simp - let ?S = "S\measure := ereal\distribution X\" - let ?T = "T\measure := ereal\distribution Y\" - interpret X: prob_space ?S by (rule distribution_prob_space) fact - interpret Y: prob_space ?T by (rule distribution_prob_space) fact - interpret XY: pair_prob_space ?S ?T by default +lemma pair_measure_eqI: + assumes "sigma_finite_measure M1" "sigma_finite_measure M2" + assumes sets: "sets (M1 \\<^isub>M M2) = sets M" + assumes emeasure: "\A B. A \ sets M1 \ B \ sets M2 \ emeasure M1 A * emeasure M2 B = emeasure M (A \ B)" + shows "M1 \\<^isub>M M2 = M" +proof - + interpret M1: sigma_finite_measure M1 by fact + interpret M2: sigma_finite_measure M2 by fact + interpret pair_sigma_finite M1 M2 by default + from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this + let ?E = "{a \ b |a b. a \ sets M1 \ b \ sets M2}" + let ?P = "M1 \\<^isub>M M2" + show ?thesis + proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) + show "?E \ Pow (space ?P)" + using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure) + show "sets ?P = sigma_sets (space ?P) ?E" + by (simp add: sets_pair_measure space_pair_measure) + then show "sets M = sigma_sets (space ?P) ?E" + using sets[symmetric] by simp + next + show "range F \ ?E" "incseq F" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" + using F by (auto simp: space_pair_measure) + next + fix X assume "X \ ?E" + then obtain A B where X[simp]: "X = A \ B" and A: "A \ sets M1" and B: "B \ sets M2" by auto + then have "emeasure ?P X = emeasure M1 A * emeasure M2 B" + by (simp add: emeasure_pair_measure_Times) + also have "\ = emeasure M (A \ B)" + using A B emeasure by auto + finally show "emeasure ?P X = emeasure M X" + by simp + qed +qed - let ?J = "XY.P\ measure := ereal \ joint_distribution X Y \" - interpret J: prob_space ?J - by (rule joint_distribution_prob_space) (simp_all add: rvs) +lemma pair_measure_eq_distr_PiM: + fixes M1 :: "'a measure" and M2 :: "'a measure" + assumes "sigma_finite_measure M1" "sigma_finite_measure M2" + shows "(M1 \\<^isub>M M2) = distr (Pi\<^isub>M UNIV (bool_case M1 M2)) (M1 \\<^isub>M M2) (\x. (x True, x False))" + (is "?P = ?D") +proof (rule pair_measure_eqI[OF assms]) + interpret B: product_sigma_finite "bool_case M1 M2" + unfolding product_sigma_finite_def using assms by (auto split: bool.split) + let ?B = "Pi\<^isub>M UNIV (bool_case M1 M2)" - have "finite_measure.\' (XY.P\ measure := ereal \ joint_distribution X Y \) A = XY.\' A" - proof (rule prob_space_unique_Int_stable) - show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P") - by fact - show "space ?P \ sets ?P" - unfolding space_pair_measure[simplified pair_measure_def space_sigma] - using X.top Y.top by (auto intro!: pair_measure_generatorI) + have [simp]: "fst \ (\x. (x True, x False)) = (\x. x True)" "snd \ (\x. (x True, x False)) = (\x. x False)" + by auto + fix A B assume A: "A \ sets M1" and B: "B \ sets M2" + have "emeasure M1 A * emeasure M2 B = (\ i\UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))" + by (simp add: UNIV_bool ac_simps) + also have "\ = emeasure ?B (Pi\<^isub>E UNIV (bool_case A B))" + using A B by (subst B.emeasure_PiM) (auto split: bool.split) + also have "Pi\<^isub>E UNIV (bool_case A B) = (\x. (x True, x False)) -` (A \ B) \ space ?B" + using A[THEN sets_into_space] B[THEN sets_into_space] + by (auto simp: Pi_iff all_bool_eq space_PiM split: bool.split) + finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \ B)" + using A B + measurable_component_singleton[of True UNIV "bool_case M1 M2"] + measurable_component_singleton[of False UNIV "bool_case M1 M2"] + by (subst emeasure_distr) (auto simp: measurable_pair_iff) +qed simp - show "prob_space ?J" by unfold_locales - show "space ?J = space ?P" - by (simp add: pair_measure_generator_def space_pair_measure) - show "sets ?J = sets (sigma ?P)" - by (simp add: pair_measure_def) +lemma measurable_Pair: + assumes rvs: "X \ measurable M S" "Y \ measurable M T" + shows "(\x. (X x, Y x)) \ measurable M (S \\<^isub>M T)" +proof - + have [simp]: "fst \ (\x. (X x, Y x)) = (\x. X x)" "snd \ (\x. (X x, Y x)) = (\x. Y x)" + by auto + show " (\x. (X x, Y x)) \ measurable M (S \\<^isub>M T)" + by (auto simp: measurable_pair_iff rvs) +qed + +lemma (in prob_space) indep_var_distribution_eq: + "indep_var S X T Y \ random_variable S X \ random_variable T Y \ + distr M S X \\<^isub>M distr M T Y = distr M (S \\<^isub>M T) (\x. (X x, Y x))" (is "_ \ _ \ _ \ ?S \\<^isub>M ?T = ?J") +proof safe + assume "indep_var S X T Y" + then show rvs: "random_variable S X" "random_variable T Y" + by (blast dest: indep_var_rv1 indep_var_rv2)+ + then have XY: "random_variable (S \\<^isub>M T) (\x. (X x, Y x))" + by (rule measurable_Pair) + + interpret X: prob_space ?S by (rule prob_space_distr) fact + interpret Y: prob_space ?T by (rule prob_space_distr) fact + interpret XY: pair_prob_space ?S ?T .. + show "?S \\<^isub>M ?T = ?J" + proof (rule pair_measure_eqI) + show "sigma_finite_measure ?S" .. + show "sigma_finite_measure ?T" .. - show "prob_space XY.P" by unfold_locales - show "space XY.P = space ?P" "sets XY.P = sets (sigma ?P)" - by (simp_all add: pair_measure_generator_def pair_measure_def) + fix A B assume A: "A \ sets ?S" and B: "B \ sets ?T" + have "emeasure ?J (A \ B) = emeasure M ((\x. (X x, Y x)) -` (A \ B) \ space M)" + using A B by (intro emeasure_distr[OF XY]) auto + also have "\ = emeasure M (X -` A \ space M) * emeasure M (Y -` B \ space M)" + using indep_varD[OF `indep_var S X T Y`, of A B] A B by (simp add: emeasure_eq_measure) + also have "\ = emeasure ?S A * emeasure ?T B" + using rvs A B by (simp add: emeasure_distr) + finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \ B)" by simp + qed simp +next + assume rvs: "random_variable S X" "random_variable T Y" + then have XY: "random_variable (S \\<^isub>M T) (\x. (X x, Y x))" + by (rule measurable_Pair) - show "A \ sets (sigma ?P)" - using `A \ sets P` unfolding P_def pair_measure_def by simp + let ?S = "distr M S X" and ?T = "distr M T Y" + interpret X: prob_space ?S by (rule prob_space_distr) fact + interpret Y: prob_space ?T by (rule prob_space_distr) fact + interpret XY: pair_prob_space ?S ?T .. + + assume "?S \\<^isub>M ?T = ?J" - fix X assume "X \ sets ?P" - then obtain A B where "A \ sets S" "B \ sets T" "X = A \ B" - by (auto simp: sets_pair_measure_generator) - then show "J.\' X = XY.\' X" - unfolding J.\'_def XY.\'_def using indep - by (simp add: XY.pair_measure_times) - (simp add: distribution_def indep_varD) + { fix S and X + have "Int_stable {X -` A \ space M |A. A \ sets S}" + proof (safe intro!: Int_stableI) + fix A B assume "A \ sets S" "B \ sets S" + then show "\C. (X -` A \ space M) \ (X -` B \ space M) = (X -` C \ space M) \ C \ sets S" + by (intro exI[of _ "A \ B"]) auto + qed } + note Int_stable = this + + show "indep_var S X T Y" unfolding indep_var_eq + proof (intro conjI indep_set_sigma_sets Int_stable rvs) + show "indep_set {X -` A \ space M |A. A \ sets S} {Y -` A \ space M |A. A \ sets T}" + proof (safe intro!: indep_setI) + { fix A assume "A \ sets S" then show "X -` A \ space M \ sets M" + using `X \ measurable M S` by (auto intro: measurable_sets) } + { fix A assume "A \ sets T" then show "Y -` A \ space M \ sets M" + using `Y \ measurable M T` by (auto intro: measurable_sets) } + next + fix A B assume ab: "A \ sets S" "B \ sets T" + then have "ereal (prob ((X -` A \ space M) \ (Y -` B \ space M))) = emeasure ?J (A \ B)" + using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"]) + also have "\ = emeasure (?S \\<^isub>M ?T) (A \ B)" + unfolding `?S \\<^isub>M ?T = ?J` .. + also have "\ = emeasure ?S A * emeasure ?T B" + using ab by (simp add: XY.emeasure_pair_measure_Times) + finally show "prob ((X -` A \ space M) \ (Y -` B \ space M)) = + prob (X -` A \ space M) * prob (Y -` B \ space M)" + using rvs ab by (simp add: emeasure_eq_measure emeasure_distr) + qed qed - then show ?thesis - using `A \ sets P` unfolding P_def J.\'_def XY.\'_def by simp qed end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Infinite_Product_Measure.thy --- a/src/HOL/Probability/Infinite_Product_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -5,9 +5,49 @@ header {*Infinite Product Measure*} theory Infinite_Product_Measure - imports Probability_Measure + imports Probability_Measure Caratheodory begin +lemma sigma_sets_mono: assumes "A \ sigma_sets X 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 \ sigma_sets X B`, auto intro: sigma_sets.intros) +qed + +lemma sigma_sets_mono': 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 sigma_sets_superset_generator: "A \ sigma_sets X A" + by (auto intro: sigma_sets.Basic) + +lemma (in product_sigma_finite) + assumes IJ: "I \ J = {}" "finite I" "finite J" and A: "A \ sets (Pi\<^isub>M (I \ J) M)" + shows emeasure_fold_integral: + "emeasure (Pi\<^isub>M (I \ J) M) A = (\\<^isup>+x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M)) \Pi\<^isub>M I M)" (is ?I) + and emeasure_fold_measurable: + "(\x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M))) \ borel_measurable (Pi\<^isub>M I M)" (is ?B) +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" .. + have merge: "(\(x, y). merge I x J y) -` A \ space (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) \ sets (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)" + by (intro measurable_sets[OF _ A] measurable_merge assms) + + show ?I + apply (subst distr_merge[symmetric, OF IJ]) + apply (subst emeasure_distr[OF measurable_merge[OF IJ(1)] A]) + apply (subst IJ.emeasure_pair_measure_alt[OF merge]) + apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure) + done + + show ?B + using IJ.measurable_emeasure_Pair1[OF merge] + by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong) +qed + lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" unfolding restrict_def extensional_def by auto @@ -41,189 +81,178 @@ qed qed -lemma (in product_prob_space) measure_preserving_restrict: +lemma prod_algebraI_finite: + "finite I \ (\i\I. E i \ sets (M i)) \ (Pi\<^isub>E I E) \ prod_algebra I M" + using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets_into_space] by simp + +lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \i\I. E i \ sets (M i)}" +proof (safe intro!: Int_stableI) + fix E F assume "\i\I. E i \ sets (M i)" "\i\I. F i \ sets (M i)" + then show "\G. Pi\<^isub>E J E \ Pi\<^isub>E J F = Pi\<^isub>E J G \ (\i\I. G i \ sets (M i))" + by (auto intro!: exI[of _ "\i. E i \ F i"]) +qed + +lemma prod_emb_trans[simp]: + "J \ K \ K \ L \ prod_emb L M K (prod_emb K M J X) = prod_emb L M J X" + by (auto simp add: Int_absorb1 prod_emb_def) + +lemma prod_emb_Pi: + assumes "X \ (\ j\J. sets (M j))" "J \ K" + shows "prod_emb K M J (Pi\<^isub>E J X) = (\\<^isub>E i\K. if i \ J then X i else space (M i))" + using assms space_closed + by (auto simp: prod_emb_def Pi_iff split: split_if_asm) blast+ + +lemma prod_emb_id: + "B \ (\\<^isub>E i\L. space (M i)) \ prod_emb L M L B = B" + by (auto simp: prod_emb_def Pi_iff subset_eq extensional_restrict) + +lemma measurable_prod_emb[intro, simp]: + "J \ L \ X \ sets (Pi\<^isub>M J M) \ prod_emb L M J X \ sets (Pi\<^isub>M L M)" + unfolding prod_emb_def space_PiM[symmetric] + by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton) + +lemma measurable_restrict_subset: "J \ L \ (\f. restrict f J) \ measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)" + by (intro measurable_restrict measurable_component_singleton) auto + +lemma (in product_prob_space) distr_restrict: assumes "J \ {}" "J \ K" "finite K" - shows "(\f. restrict f J) \ measure_preserving (\\<^isub>M i\K. M i) (\\<^isub>M i\J. M i)" (is "?R \ _") -proof - - interpret K: finite_product_prob_space M K by default fact - have J: "J \ {}" "finite J" using assms by (auto simp add: finite_subset) - interpret J: finite_product_prob_space M J - by default (insert J, auto) - from J.sigma_finite_pairs guess F .. note F = this - then have [simp,intro]: "\k i. k \ J \ F k i \ sets (M k)" - by auto - let ?F = "\i. \\<^isub>E k\J. F k i" - let ?J = "product_algebra_generator J M \ measure := measure (Pi\<^isub>M J M) \" - have "?R \ measure_preserving (\\<^isub>M i\K. M i) (sigma ?J)" - proof (rule K.measure_preserving_Int_stable) - show "Int_stable ?J" - by (auto simp: Int_stable_def product_algebra_generator_def PiE_Int) - show "range ?F \ sets ?J" "incseq ?F" "(\i. ?F i) = space ?J" - using F by auto - show "\i. measure ?J (?F i) \ \" - using F by (simp add: J.measure_times setprod_PInf) - have "measure_space (Pi\<^isub>M J M)" by default - then show "measure_space (sigma ?J)" - by (simp add: product_algebra_def sigma_def) - show "?R \ measure_preserving (Pi\<^isub>M K M) ?J" - proof (simp add: measure_preserving_def measurable_def product_algebra_generator_def del: vimage_Int, - safe intro!: restrict_extensional) - fix x k assume "k \ J" "x \ (\ i\K. space (M i))" - then show "x k \ space (M k)" using `J \ K` by auto - next - fix E assume "E \ (\ i\J. sets (M i))" - then have E: "\j. j \ J \ E j \ sets (M j)" by auto - then have *: "?R -` Pi\<^isub>E J E \ (\\<^isub>E i\K. space (M i)) = (\\<^isub>E i\K. if i \ J then E i else space (M i))" - (is "?X = Pi\<^isub>E K ?M") - using `J \ K` sets_into_space by (auto simp: Pi_iff split: split_if_asm) blast+ - with E show "?X \ sets (Pi\<^isub>M K M)" - by (auto intro!: product_algebra_generatorI) - have "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = (\i\J. measure (M i) (?M i))" - using E by (simp add: J.measure_times) - also have "\ = measure (Pi\<^isub>M K M) ?X" - unfolding * using E `finite K` `J \ K` - by (auto simp: K.measure_times M.measure_space_1 - cong del: setprod_cong - intro!: setprod_mono_one_left) - finally show "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = measure (Pi\<^isub>M K M) ?X" . - qed - qed - then show ?thesis - by (simp add: product_algebra_def sigma_def) + shows "(\\<^isub>M i\J. M i) = distr (\\<^isub>M i\K. M i) (\\<^isub>M i\J. M i) (\f. restrict f J)" (is "?P = ?D") +proof (rule measure_eqI_generator_eq) + have "finite J" using `J \ K` `finite K` by (auto simp add: finite_subset) + interpret J: finite_product_prob_space M J proof qed fact + interpret K: finite_product_prob_space M K proof qed fact + + let ?J = "{Pi\<^isub>E J E | E. \i\J. E i \ sets (M i)}" + let ?F = "\i. \\<^isub>E k\J. space (M k)" + let ?\ = "(\\<^isub>E k\J. space (M k))" + show "Int_stable ?J" + by (rule Int_stable_PiE) + show "range ?F \ ?J" "incseq ?F" "(\i. ?F i) = ?\" + using `finite J` by (auto intro!: prod_algebraI_finite) + { fix i show "emeasure ?P (?F i) \ \" by simp } + show "?J \ Pow ?\" by (auto simp: Pi_iff dest: sets_into_space) + show "sets (\\<^isub>M i\J. M i) = sigma_sets ?\ ?J" "sets ?D = sigma_sets ?\ ?J" + using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff) + + fix X assume "X \ ?J" + then obtain E where [simp]: "X = Pi\<^isub>E J E" and E: "\i\J. E i \ sets (M i)" by auto + with `finite J` have X: "X \ sets (Pi\<^isub>M J M)" by auto + + have "emeasure ?P X = (\ i\J. emeasure (M i) (E i))" + using E by (simp add: J.measure_times) + also have "\ = (\ i\J. emeasure (M i) (if i \ J then E i else space (M i)))" + by simp + also have "\ = (\ i\K. emeasure (M i) (if i \ J then E i else space (M i)))" + using `finite K` `J \ K` + by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1) + also have "\ = emeasure (Pi\<^isub>M K M) (\\<^isub>E i\K. if i \ J then E i else space (M i))" + using E by (simp add: K.measure_times) + also have "(\\<^isub>E i\K. if i \ J then E i else space (M i)) = (\f. restrict f J) -` Pi\<^isub>E J E \ (\\<^isub>E i\K. space (M i))" + using `J \ K` sets_into_space E by (force simp: Pi_iff split: split_if_asm) + finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X" + using X `J \ K` apply (subst emeasure_distr) + by (auto intro!: measurable_restrict_subset simp: space_PiM) qed -lemma (in product_prob_space) measurable_restrict: - assumes *: "J \ {}" "J \ K" "finite K" - shows "(\f. restrict f J) \ measurable (\\<^isub>M i\K. M i) (\\<^isub>M i\J. M i)" - using measure_preserving_restrict[OF *] - by (rule measure_preservingD2) +abbreviation (in product_prob_space) + "emb L K X \ prod_emb L M K X" + +lemma (in product_prob_space) emeasure_prod_emb[simp]: + assumes L: "J \ {}" "J \ L" "finite L" and X: "X \ sets (Pi\<^isub>M J M)" + shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X" + by (subst distr_restrict[OF L]) + (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X) -definition (in product_prob_space) - "emb J K X = (\x. restrict x K) -` X \ space (Pi\<^isub>M J M)" +lemma (in product_prob_space) prod_emb_injective: + assumes "J \ {}" "J \ L" "finite J" and sets: "X \ sets (Pi\<^isub>M J M)" "Y \ sets (Pi\<^isub>M J M)" + assumes "prod_emb L M J X = prod_emb L M J Y" + shows "X = Y" +proof (rule injective_vimage_restrict) + show "X \ (\\<^isub>E i\J. space (M i))" "Y \ (\\<^isub>E i\J. space (M i))" + using sets[THEN sets_into_space] by (auto simp: space_PiM) + have "\i\L. \x. x \ space (M i)" + using M.not_empty by auto + from bchoice[OF this] + show "(\\<^isub>E i\L. space (M i)) \ {}" by auto + show "(\x. restrict x J) -` X \ (\\<^isub>E i\L. space (M i)) = (\x. restrict x J) -` Y \ (\\<^isub>E i\L. space (M i))" + using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def) +qed fact -lemma (in product_prob_space) emb_trans[simp]: - "J \ K \ K \ L \ emb L K (emb K J X) = emb L J X" - by (auto simp add: Int_absorb1 emb_def) - -lemma (in product_prob_space) emb_empty[simp]: - "emb K J {} = {}" - by (simp add: emb_def) +definition (in product_prob_space) generator :: "('i \ 'a) set set" where + "generator = (\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` sets (Pi\<^isub>M J M))" -lemma (in product_prob_space) emb_Pi: - assumes "X \ (\ j\J. sets (M j))" "J \ K" - shows "emb K J (Pi\<^isub>E J X) = (\\<^isub>E i\K. if i \ J then X i else space (M i))" - using assms space_closed - by (auto simp: emb_def Pi_iff split: split_if_asm) blast+ +lemma (in product_prob_space) generatorI': + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ emb I J X \ generator" + unfolding generator_def by auto -lemma (in product_prob_space) emb_injective: - assumes "J \ {}" "J \ L" "finite J" and sets: "X \ sets (Pi\<^isub>M J M)" "Y \ sets (Pi\<^isub>M J M)" - assumes "emb L J X = emb L J Y" - shows "X = Y" -proof - - interpret J: finite_product_sigma_finite M J by default fact - show "X = Y" - proof (rule injective_vimage_restrict) - show "X \ (\\<^isub>E i\J. space (M i))" "Y \ (\\<^isub>E i\J. space (M i))" - using J.sets_into_space sets by auto - have "\i\L. \x. x \ space (M i)" - using M.not_empty by auto - from bchoice[OF this] - show "(\\<^isub>E i\L. space (M i)) \ {}" by auto - show "(\x. restrict x J) -` X \ (\\<^isub>E i\L. space (M i)) = (\x. restrict x J) -` Y \ (\\<^isub>E i\L. space (M i))" - using `emb L J X = emb L J Y` by (simp add: emb_def) - qed fact +lemma (in product_prob_space) algebra_generator: + assumes "I \ {}" shows "algebra (\\<^isub>E i\I. space (M i)) generator" (is "algebra ?\ ?G") +proof + let ?G = generator + show "?G \ Pow ?\" + by (auto simp: generator_def prod_emb_def) + from `I \ {}` obtain i where "i \ I" by auto + then show "{} \ ?G" + by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\i. {}"] + simp: sigma_sets.Empty generator_def prod_emb_def) + from `i \ I` show "?\ \ ?G" + by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\i. space (M i))"] + simp: generator_def prod_emb_def) + fix A assume "A \ ?G" + then obtain JA XA where XA: "JA \ {}" "finite JA" "JA \ I" "XA \ sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA" + by (auto simp: generator_def) + fix B assume "B \ ?G" + then obtain JB XB where XB: "JB \ {}" "finite JB" "JB \ I" "XB \ sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB" + by (auto simp: generator_def) + let ?RA = "emb (JA \ JB) JA XA" + let ?RB = "emb (JA \ JB) JB XB" + have *: "A - B = emb I (JA \ JB) (?RA - ?RB)" "A \ B = emb I (JA \ JB) (?RA \ ?RB)" + using XA A XB B by auto + show "A - B \ ?G" "A \ B \ ?G" + unfolding * using XA XB by (safe intro!: generatorI') auto qed -lemma (in product_prob_space) emb_id: - "B \ (\\<^isub>E i\L. space (M i)) \ emb L L B = B" - by (auto simp: emb_def Pi_iff subset_eq extensional_restrict) - -lemma (in product_prob_space) emb_simps: - shows "emb L K (A \ B) = emb L K A \ emb L K B" - and "emb L K (A \ B) = emb L K A \ emb L K B" - and "emb L K (A - B) = emb L K A - emb L K B" - by (auto simp: emb_def) - -lemma (in product_prob_space) measurable_emb[intro,simp]: - assumes *: "J \ {}" "J \ L" "finite L" "X \ sets (Pi\<^isub>M J M)" - shows "emb L J X \ sets (Pi\<^isub>M L M)" - using measurable_restrict[THEN measurable_sets, OF *] by (simp add: emb_def) - -lemma (in product_prob_space) measure_emb[intro,simp]: - assumes *: "J \ {}" "J \ L" "finite L" "X \ sets (Pi\<^isub>M J M)" - shows "measure (Pi\<^isub>M L M) (emb L J X) = measure (Pi\<^isub>M J M) X" - using measure_preserving_restrict[THEN measure_preservingD, OF *] - by (simp add: emb_def) - -definition (in product_prob_space) generator :: "('i \ 'a) measure_space" where - "generator = \ - space = (\\<^isub>E i\I. space (M i)), - sets = (\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` sets (Pi\<^isub>M J M)), - measure = undefined - \" +lemma (in product_prob_space) sets_PiM_generator: + assumes "I \ {}" shows "sets (PiM I M) = sigma_sets (\\<^isub>E i\I. space (M i)) generator" +proof + show "sets (Pi\<^isub>M I M) \ sigma_sets (\\<^isub>E i\I. space (M i)) generator" + unfolding sets_PiM + proof (safe intro!: sigma_sets_subseteq) + fix A assume "A \ prod_algebra I M" with `I \ {}` show "A \ generator" + by (auto intro!: generatorI' elim!: prod_algebraE) + qed +qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset) lemma (in product_prob_space) generatorI: - "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ A \ sets generator" - unfolding generator_def by auto - -lemma (in product_prob_space) generatorI': - "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ emb I J X \ sets generator" + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ A \ generator" unfolding generator_def by auto -lemma (in product_sigma_finite) - assumes "I \ J = {}" "finite I" "finite J" and A: "A \ sets (Pi\<^isub>M (I \ J) M)" - shows measure_fold_integral: - "measure (Pi\<^isub>M (I \ J) M) A = (\\<^isup>+x. measure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M)) \Pi\<^isub>M I M)" (is ?I) - and measure_fold_measurable: - "(\x. measure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M))) \ borel_measurable (Pi\<^isub>M I M)" (is ?B) -proof - - interpret I: finite_product_sigma_finite M I by default fact - interpret J: finite_product_sigma_finite M J by default fact - interpret IJ: pair_sigma_finite I.P J.P .. - show ?I - unfolding measure_fold[OF assms] - apply (subst IJ.pair_measure_alt) - apply (intro measurable_sets[OF _ A] measurable_merge assms) - apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure - intro!: I.positive_integral_cong) - done - - have "(\(x, y). merge I x J y) -` A \ space (I.P \\<^isub>M J.P) \ sets (I.P \\<^isub>M J.P)" - by (intro measurable_sets[OF _ A] measurable_merge assms) - from IJ.measure_cut_measurable_fst[OF this] - show ?B - apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure) - apply (subst (asm) measurable_cong) - apply auto - done -qed - definition (in product_prob_space) "\G A = - (THE x. \J. J \ {} \ finite J \ J \ I \ (\X\sets (Pi\<^isub>M J M). A = emb I J X \ x = measure (Pi\<^isub>M J M) X))" + (THE x. \J. J \ {} \ finite J \ J \ I \ (\X\sets (Pi\<^isub>M J M). A = emb I J X \ x = emeasure (Pi\<^isub>M J M) X))" lemma (in product_prob_space) \G_spec: assumes J: "J \ {}" "finite J" "J \ I" "A = emb I J X" "X \ sets (Pi\<^isub>M J M)" - shows "\G A = measure (Pi\<^isub>M J M) X" + shows "\G A = emeasure (Pi\<^isub>M J M) X" unfolding \G_def proof (intro the_equality allI impI ballI) fix K Y assume K: "K \ {}" "finite K" "K \ I" "A = emb I K Y" "Y \ sets (Pi\<^isub>M K M)" - have "measure (Pi\<^isub>M K M) Y = measure (Pi\<^isub>M (K \ J) M) (emb (K \ J) K Y)" + have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \ J) M) (emb (K \ J) K Y)" using K J by simp also have "emb (K \ J) K Y = emb (K \ J) J X" - using K J by (simp add: emb_injective[of "K \ J" I]) - also have "measure (Pi\<^isub>M (K \ J) M) (emb (K \ J) J X) = measure (Pi\<^isub>M J M) X" + using K J by (simp add: prod_emb_injective[of "K \ J" I]) + also have "emeasure (Pi\<^isub>M (K \ J) M) (emb (K \ J) J X) = emeasure (Pi\<^isub>M J M) X" using K J by simp - finally show "measure (Pi\<^isub>M J M) X = measure (Pi\<^isub>M K M) Y" .. + finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" .. qed (insert J, force) lemma (in product_prob_space) \G_eq: - "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ \G (emb I J X) = measure (Pi\<^isub>M J M) X" + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ \G (emb I J X) = emeasure (Pi\<^isub>M J M) X" by (intro \G_spec) auto lemma (in product_prob_space) generator_Ex: - assumes *: "A \ sets generator" - shows "\J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ \G A = measure (Pi\<^isub>M J M) X" + assumes *: "A \ generator" + shows "\J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ \G A = emeasure (Pi\<^isub>M J M) X" proof - from * obtain J X where J: "J \ {}" "finite J" "J \ I" "A = emb I J X" "X \ sets (Pi\<^isub>M J M)" unfolding generator_def by auto @@ -231,11 +260,11 @@ qed lemma (in product_prob_space) generatorE: - assumes A: "A \ sets generator" - obtains J X where "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" "emb I J X = A" "\G A = measure (Pi\<^isub>M J M) X" + assumes A: "A \ generator" + obtains J X where "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" "emb I J X = A" "\G A = emeasure (Pi\<^isub>M J M) X" proof - from generator_Ex[OF A] obtain X J where "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" "emb I J X = A" - "\G A = measure (Pi\<^isub>M J M) X" by auto + "\G A = emeasure (Pi\<^isub>M J M) X" by auto then show thesis by (intro that) auto qed @@ -243,11 +272,7 @@ assumes "finite J" "finite K" "J \ K = {}" and A: "A \ sets (Pi\<^isub>M (J \ K) M)" and x: "x \ space (Pi\<^isub>M J M)" shows "merge J x K -` A \ space (Pi\<^isub>M K M) \ sets (Pi\<^isub>M K M)" proof - - interpret J: finite_product_sigma_algebra M J by default fact - interpret K: finite_product_sigma_algebra M K by default fact - interpret JK: pair_sigma_algebra J.P K.P .. - - from JK.measurable_cut_fst[OF + from sets_Pair1[OF measurable_merge[THEN measurable_sets, OF `J \ K = {}`], OF A, of x] x show ?thesis by (simp add: space_pair_measure comp_def vimage_compose[symmetric]) @@ -266,75 +291,27 @@ have [simp]: "(K - J) \ (K \ J) = K - J" by auto have [simp]: "(K - J) \ K = K - J" by auto from y `K \ I` `J \ I` show ?thesis - by (simp split: split_merge add: emb_def Pi_iff extensional_merge_sub set_eq_iff) auto -qed - -definition (in product_prob_space) infprod_algebra :: "('i \ 'a) measure_space" where - "infprod_algebra = sigma generator \ measure := - (SOME \. (\s\sets generator. \ s = \G s) \ - prob_space \space = space generator, sets = sets (sigma generator), measure = \\)\" - -syntax - "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3PIP _:_./ _)" 10) - -syntax (xsymbols) - "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3\\<^isub>P _\_./ _)" 10) - -syntax (HTML output) - "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3\\<^isub>P _\_./ _)" 10) - -abbreviation - "Pi\<^isub>P I M \ product_prob_space.infprod_algebra M I" - -translations - "PIP x:I. M" == "CONST Pi\<^isub>P I (%x. M)" - -lemma (in product_prob_space) algebra_generator: - assumes "I \ {}" shows "algebra generator" -proof - let ?G = generator - show "sets ?G \ Pow (space ?G)" - by (auto simp: generator_def emb_def) - from `I \ {}` obtain i where "i \ I" by auto - then show "{} \ sets ?G" - by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\i. {}"] - simp: product_algebra_def sigma_def sigma_sets.Empty generator_def emb_def) - from `i \ I` show "space ?G \ sets ?G" - by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\i. space (M i))"] - simp: generator_def emb_def) - fix A assume "A \ sets ?G" - then obtain JA XA where XA: "JA \ {}" "finite JA" "JA \ I" "XA \ sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA" - by (auto simp: generator_def) - fix B assume "B \ sets ?G" - then obtain JB XB where XB: "JB \ {}" "finite JB" "JB \ I" "XB \ sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB" - by (auto simp: generator_def) - let ?RA = "emb (JA \ JB) JA XA" - let ?RB = "emb (JA \ JB) JB XB" - interpret JAB: finite_product_sigma_algebra M "JA \ JB" - by default (insert XA XB, auto) - have *: "A - B = emb I (JA \ JB) (?RA - ?RB)" "A \ B = emb I (JA \ JB) (?RA \ ?RB)" - using XA A XB B by (auto simp: emb_simps) - then show "A - B \ sets ?G" "A \ B \ sets ?G" - using XA XB by (auto intro!: generatorI') + by (simp split: split_merge add: prod_emb_def Pi_iff extensional_merge_sub set_eq_iff space_PiM) + auto qed lemma (in product_prob_space) positive_\G: assumes "I \ {}" shows "positive generator \G" proof - - interpret G!: algebra generator by (rule algebra_generator) fact + interpret G!: algebra "\\<^isub>E i\I. space (M i)" generator by (rule algebra_generator) fact show ?thesis proof (intro positive_def[THEN iffD2] conjI ballI) from generatorE[OF G.empty_sets] guess J X . note this[simp] interpret J: finite_product_sigma_finite M J by default fact have "X = {}" - by (rule emb_injective[of J I]) simp_all + by (rule prod_emb_injective[of J I]) simp_all then show "\G {} = 0" by simp next - fix A assume "A \ sets generator" + fix A assume "A \ generator" from generatorE[OF this] guess J X . note this[simp] interpret J: finite_product_sigma_finite M J by default fact - show "0 \ \G A" by simp + show "0 \ \G A" by (simp add: emeasure_nonneg) qed qed @@ -342,102 +319,47 @@ assumes "I \ {}" shows "additive generator \G" proof - - interpret G!: algebra generator by (rule algebra_generator) fact + interpret G!: algebra "\\<^isub>E i\I. space (M i)" generator by (rule algebra_generator) fact show ?thesis proof (intro additive_def[THEN iffD2] ballI impI) - fix A assume "A \ sets generator" with generatorE guess J X . note J = this - fix B assume "B \ sets generator" with generatorE guess K Y . note K = this + fix A assume "A \ generator" with generatorE guess J X . note J = this + fix B assume "B \ generator" with generatorE guess K Y . note K = this assume "A \ B = {}" have JK: "J \ K \ {}" "J \ K \ I" "finite (J \ K)" using J K by auto interpret JK: finite_product_sigma_finite M "J \ K" by default fact have JK_disj: "emb (J \ K) J X \ emb (J \ K) K Y = {}" - apply (rule emb_injective[of "J \ K" I]) + apply (rule prod_emb_injective[of "J \ K" I]) apply (insert `A \ B = {}` JK J K) - apply (simp_all add: JK.Int emb_simps) + apply (simp_all add: Int prod_emb_Int) done have AB: "A = emb I (J \ K) (emb (J \ K) J X)" "B = emb I (J \ K) (emb (J \ K) K Y)" using J K by simp_all then have "\G (A \ B) = \G (emb I (J \ K) (emb (J \ K) J X \ emb (J \ K) K Y))" - by (simp add: emb_simps) - also have "\ = measure (Pi\<^isub>M (J \ K) M) (emb (J \ K) J X \ emb (J \ K) K Y)" - using JK J(1, 4) K(1, 4) by (simp add: \G_eq JK.Un) + by simp + also have "\ = emeasure (Pi\<^isub>M (J \ K) M) (emb (J \ K) J X \ emb (J \ K) K Y)" + using JK J(1, 4) K(1, 4) by (simp add: \G_eq Un del: prod_emb_Un) also have "\ = \G A + \G B" - using J K JK_disj by (simp add: JK.measure_additive[symmetric]) + using J K JK_disj by (simp add: plus_emeasure[symmetric]) finally show "\G (A \ B) = \G A + \G B" . qed qed -lemma (in product_prob_space) finite_index_eq_finite_product: - assumes "finite I" - shows "sets (sigma generator) = sets (Pi\<^isub>M I M)" -proof safe - interpret I: finite_product_sigma_algebra M I by default fact - have space_generator[simp]: "space generator = space (Pi\<^isub>M I M)" - by (simp add: generator_def product_algebra_def) - { fix A assume "A \ sets (sigma generator)" - then show "A \ sets I.P" unfolding sets_sigma - proof induct - case (Basic A) - from generatorE[OF this] guess J X . note J = this - with `finite I` have "emb I J X \ sets I.P" by auto - with `emb I J X = A` show "A \ sets I.P" by simp - qed auto } - { fix A assume A: "A \ sets I.P" - show "A \ sets (sigma generator)" - proof cases - assume "I = {}" - with I.P_empty[OF this] A - have "A = space generator \ A = {}" - unfolding space_generator by auto - then show ?thesis - by (auto simp: sets_sigma simp del: space_generator - intro: sigma_sets.Empty sigma_sets_top) - next - assume "I \ {}" - note A this - moreover with I.sets_into_space have "emb I I A = A" by (intro emb_id) auto - ultimately show "A \ sets (sigma generator)" - using `finite I` unfolding sets_sigma - by (intro sigma_sets.Basic generatorI[of I A]) auto - qed } -qed - -lemma (in product_prob_space) extend_\G: - "\\. (\s\sets generator. \ s = \G s) \ - prob_space \space = space generator, sets = sets (sigma generator), measure = \\" +lemma (in product_prob_space) emeasure_PiM_emb_not_empty: + assumes X: "J \ {}" "J \ I" "finite J" "\i\J. X i \ sets (M i)" + shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)" proof cases - assume "finite I" - interpret I: finite_product_prob_space M I by default fact - show ?thesis - proof (intro exI[of _ "measure (Pi\<^isub>M I M)"] ballI conjI) - fix A assume "A \ sets generator" - from generatorE[OF this] guess J X . note J = this - from J(1-4) `finite I` show "measure I.P A = \G A" - unfolding J(6) - by (subst J(5)[symmetric]) (simp add: measure_emb) - next - have [simp]: "space generator = space (Pi\<^isub>M I M)" - by (simp add: generator_def product_algebra_def) - have "\space = space generator, sets = sets (sigma generator), measure = measure I.P\ - = I.P" (is "?P = _") - by (auto intro!: measure_space.equality simp: finite_index_eq_finite_product[OF `finite I`]) - show "prob_space ?P" - proof - show "measure_space ?P" using `?P = I.P` by simp default - show "measure ?P (space ?P) = 1" - using I.measure_space_1 by simp - qed - qed + assume "finite I" with X show ?thesis by simp next + let ?\ = "\\<^isub>E i\I. space (M i)" let ?G = generator assume "\ finite I" then have I_not_empty: "I \ {}" by auto - interpret G!: algebra generator by (rule algebra_generator) fact + interpret G!: algebra ?\ generator by (rule algebra_generator) fact note \G_mono = G.additive_increasing[OF positive_\G[OF I_not_empty] additive_\G[OF I_not_empty], THEN increasingD] - { fix Z J assume J: "J \ {}" "finite J" "J \ I" and Z: "Z \ sets ?G" + { fix Z J assume J: "J \ {}" "finite J" "J \ I" and Z: "Z \ ?G" from `infinite I` `finite J` obtain k where k: "k \ I" "k \ J" by (metis rev_finite_subset subsetI) @@ -445,7 +367,7 @@ moreover def K \ "insert k K'" moreover def X \ "emb K K' X'" ultimately have K: "K \ {}" "finite K" "K \ I" "X \ sets (Pi\<^isub>M K M)" "Z = emb I K X" - "K - J \ {}" "K - J \ I" "\G Z = measure (Pi\<^isub>M K M) X" + "K - J \ {}" "K - J \ I" "\G Z = emeasure (Pi\<^isub>M K M) X" by (auto simp: subset_insertI) let ?M = "\y. merge J y (K - J) -` emb (J \ K) K X \ space (Pi\<^isub>M (K - J) M)" @@ -455,9 +377,9 @@ have **: "?M y \ sets (Pi\<^isub>M (K - J) M)" using J K y by (intro merge_sets) auto ultimately - have ***: "(merge J y (I - J) -` Z \ space (Pi\<^isub>M I M)) \ sets ?G" + have ***: "(merge J y (I - J) -` Z \ space (Pi\<^isub>M I M)) \ ?G" using J K by (intro generatorI) auto - have "\G (merge J y (I - J) -` emb I K X \ space (Pi\<^isub>M I M)) = measure (Pi\<^isub>M (K - J) M) (?M y)" + have "\G (merge J y (I - J) -` emb I K X \ space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)" unfolding * using K J by (subst \G_eq[OF _ _ _ **]) auto note * ** *** this } note merge_in_G = this @@ -467,16 +389,16 @@ interpret J: finite_product_prob_space M J by default fact+ interpret KmJ: finite_product_prob_space M "K - J" by default fact+ - have "\G Z = measure (Pi\<^isub>M (J \ (K - J)) M) (emb (J \ (K - J)) K X)" + have "\G Z = emeasure (Pi\<^isub>M (J \ (K - J)) M) (emb (J \ (K - J)) K X)" using K J by simp - also have "\ = (\\<^isup>+ x. measure (Pi\<^isub>M (K - J) M) (?M x) \Pi\<^isub>M J M)" - using K J by (subst measure_fold_integral) auto + also have "\ = (\\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \Pi\<^isub>M J M)" + using K J by (subst emeasure_fold_integral) auto also have "\ = (\\<^isup>+ y. \G (merge J y (I - J) -` Z \ space (Pi\<^isub>M I M)) \Pi\<^isub>M J M)" (is "_ = (\\<^isup>+x. \G (?MZ x) \Pi\<^isub>M J M)") - proof (intro J.positive_integral_cong) + proof (intro positive_integral_cong) fix x assume x: "x \ space (Pi\<^isub>M J M)" with K merge_in_G(2)[OF this] - show "measure (Pi\<^isub>M (K - J) M) (?M x) = \G (?MZ x)" + show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \G (?MZ x)" unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \G_eq) auto qed finally have fold: "\G Z = (\\<^isup>+x. \G (?MZ x) \Pi\<^isub>M J M)" . @@ -490,21 +412,18 @@ let ?q = "\y. \G (merge J y (I - J) -` Z \ space (Pi\<^isub>M I M))" have "?q \ borel_measurable (Pi\<^isub>M J M)" unfolding `Z = emb I K X` using J K merge_in_G(3) - by (simp add: merge_in_G \G_eq measure_fold_measurable - del: space_product_algebra cong: measurable_cong) + by (simp add: merge_in_G \G_eq emeasure_fold_measurable cong: measurable_cong) note this fold le_1 merge_in_G(3) } note fold = this - have "\\. (\s\sets ?G. \ s = \G s) \ - measure_space \space = space ?G, sets = sets (sigma ?G), measure = \\" - (is "\\. _ \ measure_space (?ms \)") + have "\\. (\s\?G. \ s = \G s) \ measure_space ?\ (sigma_sets ?\ ?G) \" proof (rule G.caratheodory_empty_continuous[OF positive_\G additive_\G]) - fix A assume "A \ sets ?G" + fix A assume "A \ ?G" with generatorE guess J X . note JX = this interpret JK: finite_product_prob_space M J by default fact+ from JX show "\G A \ \" by simp next - fix A assume A: "range A \ sets ?G" "decseq A" "(\i. A i) = {}" + fix A assume A: "range A \ ?G" "decseq A" "(\i. A i) = {}" then have "decseq (\i. \G (A i))" by (auto intro!: \G_mono simp: decseq_def) moreover @@ -515,7 +434,7 @@ using A positive_\G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def) ultimately have "0 < ?a" by auto - have "\n. \J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A n = emb I J X \ \G (A n) = measure (Pi\<^isub>M J M) X" + have "\n. \J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A n = emb I J X \ \G (A n) = emeasure (Pi\<^isub>M J M) X" using A by (intro allI generator_Ex) auto then obtain J' X' where J': "\n. J' n \ {}" "\n. finite (J' n)" "\n. J' n \ I" "\n. X' n \ sets (Pi\<^isub>M (J' n) M)" and A': "\n. A n = emb I (J' n) (X' n)" @@ -524,8 +443,8 @@ moreover def X \ "\n. emb (J n) (J' n) (X' n)" ultimately have J: "\n. J n \ {}" "\n. finite (J n)" "\n. J n \ I" "\n. X n \ sets (Pi\<^isub>M (J n) M)" by auto - with A' have A_eq: "\n. A n = emb I (J n) (X n)" "\n. A n \ sets ?G" - unfolding J_def X_def by (subst emb_trans) (insert A, auto) + with A' have A_eq: "\n. A n = emb I (J n) (X n)" "\n. A n \ ?G" + unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto) have J_mono: "\n m. n \ m \ J n \ J m" unfolding J_def by force @@ -538,8 +457,8 @@ let ?M = "\K Z y. merge K y (I - K) -` Z \ space (Pi\<^isub>M I M)" - { fix Z k assume Z: "range Z \ sets ?G" "decseq Z" "\n. ?a / 2^k \ \G (Z n)" - then have Z_sets: "\n. Z n \ sets ?G" by auto + { fix Z k assume Z: "range Z \ ?G" "decseq Z" "\n. ?a / 2^k \ \G (Z n)" + then have Z_sets: "\n. Z n \ ?G" by auto fix J' assume J': "J' \ {}" "finite J'" "J' \ I" interpret J': finite_product_prob_space M J' by default fact+ @@ -552,13 +471,13 @@ by (rule measurable_sets) auto } note Q_sets = this - have "?a / 2^(k+1) \ (INF n. measure (Pi\<^isub>M J' M) (?Q n))" + have "?a / 2^(k+1) \ (INF n. emeasure (Pi\<^isub>M J' M) (?Q n))" proof (intro INF_greatest) fix n have "?a / 2^k \ \G (Z n)" using Z by auto also have "\ \ (\\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \Pi\<^isub>M J' M)" - unfolding fold(2)[OF J' `Z n \ sets ?G`] - proof (intro J'.positive_integral_mono) + unfolding fold(2)[OF J' `Z n \ ?G`] + proof (intro positive_integral_mono) fix x assume x: "x \ space (Pi\<^isub>M J' M)" then have "?q n x \ 1 + 0" using J' Z fold(3) Z_sets by auto @@ -568,15 +487,15 @@ with x show "?q n x \ indicator (?Q n) x + ?a / 2^(k+1)" by (auto split: split_indicator simp del: power_Suc) qed - also have "\ = measure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)" - using `0 \ ?a` Q_sets J'.measure_space_1 - by (subst J'.positive_integral_add) auto - finally show "?a / 2^(k+1) \ measure (Pi\<^isub>M J' M) (?Q n)" using `?a \ 1` - by (cases rule: ereal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"]) + also have "\ = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)" + using `0 \ ?a` Q_sets J'.emeasure_space_1 + by (subst positive_integral_add) auto + finally show "?a / 2^(k+1) \ emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \ 1` + by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"]) (auto simp: field_simps) qed - also have "\ = measure (Pi\<^isub>M J' M) (\n. ?Q n)" - proof (intro J'.continuity_from_above) + also have "\ = emeasure (Pi\<^isub>M J' M) (\n. ?Q n)" + proof (intro INF_emeasure_decseq) show "range ?Q \ sets (Pi\<^isub>M J' M)" using Q_sets by auto show "decseq ?Q" unfolding decseq_def @@ -587,13 +506,13 @@ also have "?q n x \ ?q m x" proof (rule \G_mono) from fold(4)[OF J', OF Z_sets x] - show "?M J' (Z n) x \ sets ?G" "?M J' (Z m) x \ sets ?G" by auto + show "?M J' (Z n) x \ ?G" "?M J' (Z m) x \ ?G" by auto show "?M J' (Z n) x \ ?M J' (Z m) x" using `decseq Z`[THEN decseqD, OF `m \ n`] by auto qed finally show "?a / 2^(k+1) \ ?q m x" . qed - qed (intro J'.finite_measure Q_sets) + qed simp finally have "(\n. ?Q n) \ {}" using `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) then have "\w\space (Pi\<^isub>M J' M). \n. ?a / 2 ^ (k + 1) \ ?q n w" by auto } @@ -631,12 +550,12 @@ show "w k \ space (Pi\<^isub>M (J (Suc k)) M)" using Suc by simp then show "restrict (w k) (J k) = w k" - by (simp add: extensional_restrict) + by (simp add: extensional_restrict space_PiM) qed next assume "J k \ J (Suc k)" with J_mono[of k "Suc k"] have "J (Suc k) - J k \ {}" (is "?D \ {}") by auto - have "range (\n. ?M (J k) (A n) (w k)) \ sets ?G" + have "range (\n. ?M (J k) (A n) (w k)) \ ?G" "decseq (\n. ?M (J k) (A n) (w k))" "\n. ?a / 2 ^ (k + 1) \ \G (?M (J k) (A n) (w k))" using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc @@ -651,11 +570,11 @@ by (auto intro!: ext split: split_merge) have *: "\n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w" using w'(1) J(3)[of "Suc k"] - by (auto split: split_merge intro!: extensional_merge_sub) force+ + by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+ show ?thesis apply (rule exI[of _ ?w]) using w' J_mono[of k "Suc k"] wk unfolding * - apply (auto split: split_merge intro!: extensional_merge_sub ext) + apply (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM) apply (force simp: extensional_def) done qed @@ -675,7 +594,7 @@ then have "merge (J k) (w k) (I - J k) x \ A k" by auto then have "\x\A k. restrict x (J k) = w k" using `w k \ space (Pi\<^isub>M (J k) M)` - by (intro rev_bexI) (auto intro!: ext simp: extensional_def) + by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) ultimately have "w k \ space (Pi\<^isub>M (J k) M)" "\x\A k. restrict x (J k) = w k" "k \ 0 \ restrict (w k) (J (k - 1)) = w (k - 1)" @@ -707,17 +626,17 @@ have w'_simps2: "\i. i \ (\k. J k) \ i \ I \ w' i \ space (M i)" using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]]) { fix i assume "i \ I" then have "w' i \ space (M i)" - using w(1) by (cases "i \ (\k. J k)") (force simp: w'_simps2 w'_eq)+ } + using w(1) by (cases "i \ (\k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ } note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this have w': "w' \ space (Pi\<^isub>M I M)" - using w(1) by (auto simp add: Pi_iff extensional_def) + using w(1) by (auto simp add: Pi_iff extensional_def space_PiM) { fix n have "restrict w' (J n) = w n" using w(1) - by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def) + by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def space_PiM) with w[of n] obtain x where "x \ A n" "restrict x (J n) = restrict w' (J n)" by auto - then have "w' \ A n" unfolding A_eq using w' by (auto simp: emb_def) } + then have "w' \ A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) } then have "w' \ (\i. A i)" by auto with `(\i. A i) = {}` show False by auto qed @@ -726,276 +645,76 @@ qed fact+ then guess \ .. note \ = this show ?thesis - proof (intro exI[of _ \] conjI) - show "\S\sets ?G. \ S = \G S" using \ by simp - show "prob_space (?ms \)" - proof - show "measure_space (?ms \)" using \ by simp - obtain i where "i \ I" using I_not_empty by auto - interpret i: finite_product_sigma_finite M "{i}" by default auto - let ?X = "\\<^isub>E i\{i}. space (M i)" - have X: "?X \ sets (Pi\<^isub>M {i} M)" - by auto - with `i \ I` have "emb I {i} ?X \ sets generator" - by (intro generatorI') auto - with \ have "\ (emb I {i} ?X) = \G (emb I {i} ?X)" by auto - with \G_eq[OF _ _ _ X] `i \ I` - have "\ (emb I {i} ?X) = measure (M i) (space (M i))" - by (simp add: i.measure_times) - also have "emb I {i} ?X = space (Pi\<^isub>P I M)" - using `i \ I` by (auto simp: emb_def infprod_algebra_def generator_def) - finally show "measure (?ms \) (space (?ms \)) = 1" - using M.measure_space_1 by (simp add: infprod_algebra_def) - qed + proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \ J X]) + from assms show "(J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))" + by (simp add: Pi_iff) + next + fix J X assume J: "(J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))" + then show "emb I J (Pi\<^isub>E J X) \ Pow (\\<^isub>E i\I. space (M i))" + by (auto simp: Pi_iff prod_emb_def dest: sets_into_space) + have "emb I J (Pi\<^isub>E J X) \ generator" + using J `I \ {}` by (intro generatorI') auto + then have "\ (emb I J (Pi\<^isub>E J X)) = \G (emb I J (Pi\<^isub>E J X))" + using \ by simp + also have "\ = (\ j\J. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" + using J `I \ {}` by (subst \G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff) + also have "\ = (\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. + if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" + using J `I \ {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) + finally show "\ (emb I J (Pi\<^isub>E J X)) = \" . + next + let ?F = "\j. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j))" + have "(\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. ?F j) = (\j\J. ?F j)" + using X `I \ {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1) + then show "(\j\J \ {i \ I. emeasure (M i) (space (M i)) \ 1}. ?F j) = + emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)" + using X by (auto simp add: emeasure_PiM) + next + show "positive (sets (Pi\<^isub>M I M)) \" "countably_additive (sets (Pi\<^isub>M I M)) \" + using \ unfolding sets_PiM_generator[OF `I \ {}`] by (auto simp: measure_space_def) qed qed -lemma (in product_prob_space) infprod_spec: - "(\s\sets generator. measure (Pi\<^isub>P I M) s = \G s) \ prob_space (Pi\<^isub>P I M)" - (is "?Q infprod_algebra") - unfolding infprod_algebra_def - by (rule someI2_ex[OF extend_\G]) - (auto simp: sigma_def generator_def) - -sublocale product_prob_space \ P: prob_space "Pi\<^isub>P I M" - using infprod_spec by simp - -lemma (in product_prob_space) measure_infprod_emb: - assumes "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" - shows "\ (emb I J X) = measure (Pi\<^isub>M J M) X" -proof - - have "emb I J X \ sets generator" - using assms by (rule generatorI') - with \G_eq[OF assms] infprod_spec show ?thesis by auto -qed - -lemma (in product_prob_space) measurable_component: - assumes "i \ I" - shows "(\x. x i) \ measurable (Pi\<^isub>P I M) (M i)" -proof (unfold measurable_def, safe) - fix x assume "x \ space (Pi\<^isub>P I M)" - then show "x i \ space (M i)" - using `i \ I` by (auto simp: infprod_algebra_def generator_def) -next - fix A assume "A \ sets (M i)" - with `i \ I` have - "(\\<^isub>E x \ {i}. A) \ sets (Pi\<^isub>M {i} M)" - "(\x. x i) -` A \ space (Pi\<^isub>P I M) = emb I {i} (\\<^isub>E x \ {i}. A)" - by (auto simp: infprod_algebra_def generator_def emb_def) - from generatorI[OF _ _ _ this] `i \ I` - show "(\x. x i) -` A \ space (Pi\<^isub>P I M) \ sets (Pi\<^isub>P I M)" - unfolding infprod_algebra_def by auto -qed - -lemma (in product_prob_space) emb_in_infprod_algebra[intro]: - fixes J assumes J: "finite J" "J \ I" and X: "X \ sets (Pi\<^isub>M J M)" - shows "emb I J X \ sets (\\<^isub>P i\I. M i)" -proof cases - assume "J = {}" - with X have "emb I J X = space (\\<^isub>P i\I. M i) \ emb I J X = {}" - by (auto simp: emb_def infprod_algebra_def generator_def - product_algebra_def product_algebra_generator_def image_constant sigma_def) - then show ?thesis by auto -next - assume "J \ {}" - show ?thesis unfolding infprod_algebra_def - by simp (intro in_sigma generatorI' `J \ {}` J X) -qed - -lemma (in product_prob_space) finite_measure_infprod_emb: - assumes "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" - shows "\' (emb I J X) = finite_measure.\' (Pi\<^isub>M J M) X" -proof - - interpret J: finite_product_prob_space M J by default fact+ - from assms have "emb I J X \ sets (Pi\<^isub>P I M)" by auto - with assms show "\' (emb I J X) = J.\' X" - unfolding \'_def J.\'_def - unfolding measure_infprod_emb[OF assms] - by auto -qed - -lemma (in finite_product_prob_space) finite_measure_times: - assumes "\i. i \ I \ A i \ sets (M i)" - shows "\' (Pi\<^isub>E I A) = (\i\I. M.\' i (A i))" - using assms - unfolding \'_def M.\'_def - by (subst measure_times[OF assms]) - (auto simp: finite_measure_eq M.finite_measure_eq setprod_ereal) - -lemma (in product_prob_space) finite_measure_infprod_emb_Pi: - assumes J: "finite J" "J \ I" "\j. j \ J \ X j \ sets (M j)" - shows "\' (emb I J (Pi\<^isub>E J X)) = (\j\J. M.\' j (X j))" -proof cases - assume "J = {}" - then have "emb I J (Pi\<^isub>E J X) = space infprod_algebra" - by (auto simp: infprod_algebra_def generator_def sigma_def emb_def) - then show ?thesis using `J = {}` P.prob_space - by simp -next - assume "J \ {}" - interpret J: finite_product_prob_space M J by default fact+ - have "(\i\J. M.\' i (X i)) = J.\' (Pi\<^isub>E J X)" - using J `J \ {}` by (subst J.finite_measure_times) auto - also have "\ = \' (emb I J (Pi\<^isub>E J X))" - using J `J \ {}` by (intro finite_measure_infprod_emb[symmetric]) auto - finally show ?thesis by simp -qed - -lemma sigma_sets_mono: assumes "A \ sigma_sets X B" shows "sigma_sets X A \ sigma_sets X B" +sublocale product_prob_space \ P: prob_space "Pi\<^isub>M I M" proof - fix x assume "x \ sigma_sets X A" then show "x \ sigma_sets X B" - by induct (insert `A \ sigma_sets X B`, auto intro: sigma_sets.intros) -qed - -lemma sigma_sets_mono': 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) + show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1" + proof cases + assume "I = {}" then show ?thesis by (simp add: space_PiM_empty) + next + assume "I \ {}" + then obtain i where "i \ I" by auto + moreover then have "emb I {i} (\\<^isub>E i\{i}. space (M i)) = (space (Pi\<^isub>M I M))" + by (auto simp: prod_emb_def space_PiM) + ultimately show ?thesis + using emeasure_PiM_emb_not_empty[of "{i}" "\i. space (M i)"] + by (simp add: emeasure_PiM emeasure_space_1) + qed qed -lemma sigma_sets_superset_generator: "A \ sigma_sets X A" - by (auto intro: sigma_sets.Basic) - -lemma (in product_prob_space) infprod_algebra_alt: - "Pi\<^isub>P I M = sigma \ space = space (Pi\<^isub>P I M), - sets = (\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` Pi\<^isub>E J ` (\ i \ J. sets (M i))), - measure = measure (Pi\<^isub>P I M) \" - (is "_ = sigma \ space = ?O, sets = ?M, measure = ?m \") -proof (rule measure_space.equality) - let ?G = "\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` sets (Pi\<^isub>M J M)" - have "sigma_sets ?O ?M = sigma_sets ?O ?G" - proof (intro equalityI sigma_sets_mono UN_least) - fix J assume J: "J \ {J. J \ {} \ finite J \ J \ I}" - have "emb I J ` Pi\<^isub>E J ` (\ i\J. sets (M i)) \ emb I J ` sets (Pi\<^isub>M J M)" by auto - also have "\ \ ?G" using J by (rule UN_upper) - also have "\ \ sigma_sets ?O ?G" by (rule sigma_sets_superset_generator) - finally show "emb I J ` Pi\<^isub>E J ` (\ i\J. sets (M i)) \ sigma_sets ?O ?G" . - have "emb I J ` sets (Pi\<^isub>M J M) = emb I J ` sigma_sets (space (Pi\<^isub>M J M)) (Pi\<^isub>E J ` (\ i \ J. sets (M i)))" - by (simp add: sets_sigma product_algebra_generator_def product_algebra_def) - also have "\ = sigma_sets (space (Pi\<^isub>M I M)) (emb I J ` Pi\<^isub>E J ` (\ i \ J. sets (M i)))" - using J M.sets_into_space - by (auto simp: emb_def [abs_def] intro!: sigma_sets_vimage[symmetric]) blast - also have "\ \ sigma_sets (space (Pi\<^isub>M I M)) ?M" - using J by (intro sigma_sets_mono') auto - finally show "emb I J ` sets (Pi\<^isub>M J M) \ sigma_sets ?O ?M" - by (simp add: infprod_algebra_def generator_def) - qed - then show "sets (Pi\<^isub>P I M) = sets (sigma \ space = ?O, sets = ?M, measure = ?m \)" - by (simp_all add: infprod_algebra_def generator_def sets_sigma) -qed simp_all - -lemma (in product_prob_space) infprod_algebra_alt2: - "Pi\<^isub>P I M = sigma \ space = space (Pi\<^isub>P I M), - sets = (\i\I. (\A. (\x. x i) -` A \ space (Pi\<^isub>P I M)) ` sets (M i)), - measure = measure (Pi\<^isub>P I M) \" - (is "_ = ?S") -proof (rule measure_space.equality) - let "sigma \ space = ?O, sets = ?A, \ = _ \" = ?S - let ?G = "(\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` Pi\<^isub>E J ` (\ i \ J. sets (M i)))" - have "sets (Pi\<^isub>P I M) = sigma_sets ?O ?G" - by (subst infprod_algebra_alt) (simp add: sets_sigma) - also have "\ = sigma_sets ?O ?A" - proof (intro equalityI sigma_sets_mono subsetI) - interpret A: sigma_algebra ?S - by (rule sigma_algebra_sigma) auto - fix A assume "A \ ?G" - then obtain J B where "finite J" "J \ {}" "J \ I" "A = emb I J (Pi\<^isub>E J B)" - and B: "\i. i \ J \ B i \ sets (M i)" - by auto - then have A: "A = (\j\J. (\x. x j) -` (B j) \ space (Pi\<^isub>P I M))" - by (auto simp: emb_def infprod_algebra_def generator_def Pi_iff) - { fix j assume "j\J" - with `J \ I` have "j \ I" by auto - with `j \ J` B have "(\x. x j) -` (B j) \ space (Pi\<^isub>P I M) \ sets ?S" - by (auto simp: sets_sigma intro: sigma_sets.Basic) } - with `finite J` `J \ {}` have "A \ sets ?S" - unfolding A by (intro A.finite_INT) auto - then show "A \ sigma_sets ?O ?A" by (simp add: sets_sigma) - next - fix A assume "A \ ?A" - then obtain i B where i: "i \ I" "B \ sets (M i)" - and "A = (\x. x i) -` B \ space (Pi\<^isub>P I M)" - by auto - then have "A = emb I {i} (Pi\<^isub>E {i} (\_. B))" - by (auto simp: emb_def infprod_algebra_def generator_def Pi_iff) - with i show "A \ sigma_sets ?O ?G" - by (intro sigma_sets.Basic UN_I[where a="{i}"]) auto - qed - also have "\ = sets ?S" - by (simp add: sets_sigma) - finally show "sets (Pi\<^isub>P I M) = sets ?S" . -qed simp_all - -lemma (in product_prob_space) measurable_into_infprod_algebra: - assumes "sigma_algebra N" - assumes f: "\i. i \ I \ (\x. f x i) \ measurable N (M i)" - assumes ext: "\x. x \ space N \ f x \ extensional I" - shows "f \ measurable N (Pi\<^isub>P I M)" -proof - - interpret N: sigma_algebra N by fact - have f_in: "\i. i \ I \ (\x. f x i) \ space N \ space (M i)" - using f by (auto simp: measurable_def) - { fix i A assume i: "i \ I" "A \ sets (M i)" - then have "f -` (\x. x i) -` A \ f -` space infprod_algebra \ space N = (\x. f x i) -` A \ space N" - using f_in ext by (auto simp: infprod_algebra_def generator_def) - also have "\ \ sets N" - by (rule measurable_sets f i)+ - finally have "f -` (\x. x i) -` A \ f -` space infprod_algebra \ space N \ sets N" . } - with f_in ext show ?thesis - by (subst infprod_algebra_alt2) - (auto intro!: N.measurable_sigma simp: Pi_iff infprod_algebra_def generator_def) +lemma (in product_prob_space) emeasure_PiM_emb: + assumes X: "J \ I" "finite J" "\i. i \ J \ X i \ sets (M i)" + shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\ i\J. emeasure (M i) (X i))" +proof cases + assume "J = {}" + moreover have "emb I {} {\x. undefined} = space (Pi\<^isub>M I M)" + by (auto simp: space_PiM prod_emb_def) + ultimately show ?thesis + by (simp add: space_PiM_empty P.emeasure_space_1) +next + assume "J \ {}" with X show ?thesis + by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM) qed -lemma (in product_prob_space) measurable_singleton_infprod: - assumes "i \ I" - shows "(\x. x i) \ measurable (Pi\<^isub>P I M) (M i)" -proof (unfold measurable_def, intro CollectI conjI ballI) - show "(\x. x i) \ space (Pi\<^isub>P I M) \ space (M i)" - using M.sets_into_space `i \ I` - by (auto simp: infprod_algebra_def generator_def) - fix A assume "A \ sets (M i)" - have "(\x. x i) -` A \ space (Pi\<^isub>P I M) = emb I {i} (\\<^isub>E _\{i}. A)" - by (auto simp: infprod_algebra_def generator_def emb_def) - also have "\ \ sets (Pi\<^isub>P I M)" - using `i \ I` `A \ sets (M i)` - by (intro emb_in_infprod_algebra product_algebraI) auto - finally show "(\x. x i) -` A \ space (Pi\<^isub>P I M) \ sets (Pi\<^isub>P I M)" . -qed +lemma (in product_prob_space) measure_PiM_emb: + assumes "J \ I" "finite J" "\i. i \ J \ X i \ sets (M i)" + shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\ i\J. measure (M i) (X i))" + using emeasure_PiM_emb[OF assms] + unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal) -lemma (in product_prob_space) sigma_product_algebra_sigma_eq: - assumes M: "\i. i \ I \ M i = sigma (E i)" - shows "sets (Pi\<^isub>P I M) = sigma_sets (space (Pi\<^isub>P I M)) (\i\I. (\A. (\x. x i) -` A \ space (Pi\<^isub>P I M)) ` sets (E i))" -proof - - let ?E = "(\i\I. (\A. (\x. x i) -` A \ space (Pi\<^isub>P I M)) ` sets (E i))" - let ?M = "(\i\I. (\A. (\x. x i) -` A \ space (Pi\<^isub>P I M)) ` sets (M i))" - { fix i A assume "i\I" "A \ sets (E i)" - then have "A \ sets (M i)" using M by auto - then have "A \ Pow (space (M i))" using M.sets_into_space by auto - then have "A \ Pow (space (E i))" using M[OF `i \ I`] by auto } - moreover - have "\i. i \ I \ (\x. x i) \ space infprod_algebra \ space (E i)" - by (auto simp: M infprod_algebra_def generator_def Pi_iff) - ultimately have "sigma_sets (space (Pi\<^isub>P I M)) ?M \ sigma_sets (space (Pi\<^isub>P I M)) ?E" - apply (intro sigma_sets_mono UN_least) - apply (simp add: sets_sigma M) - apply (subst sigma_sets_vimage[symmetric]) - apply (auto intro!: sigma_sets_mono') - done - moreover have "sigma_sets (space (Pi\<^isub>P I M)) ?E \ sigma_sets (space (Pi\<^isub>P I M)) ?M" - by (intro sigma_sets_mono') (auto simp: M) - ultimately show ?thesis - by (subst infprod_algebra_alt2) (auto simp: sets_sigma) -qed - -lemma (in product_prob_space) Int_proj_eq_emb: - assumes "J \ {}" "J \ I" - shows "(\i\J. (\x. x i) -` A i \ space (Pi\<^isub>P I M)) = emb I J (Pi\<^isub>E J A)" - using assms by (auto simp: infprod_algebra_def generator_def emb_def Pi_iff) - -lemma (in product_prob_space) emb_insert: - "i \ J \ emb I J (Pi\<^isub>E J f) \ ((\x. x i) -` A \ space (Pi\<^isub>P I M)) = - emb I (insert i J) (Pi\<^isub>E (insert i J) (f(i := A)))" - by (auto simp: emb_def Pi_iff infprod_algebra_def generator_def split: split_if_asm) +lemma (in finite_product_prob_space) finite_measure_PiM_emb: + "(\i. i \ I \ A i \ sets (M i)) \ measure (PiM I M) (Pi\<^isub>E I A) = (\i\I. measure (M i) (A i))" + using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets_into_space, of I A M] + by auto subsection {* Sequence space *} @@ -1003,36 +722,30 @@ lemma (in sequence_space) infprod_in_sets[intro]: fixes E :: "nat \ 'a set" assumes E: "\i. E i \ sets (M i)" - shows "Pi UNIV E \ sets (Pi\<^isub>P UNIV M)" + shows "Pi UNIV E \ sets (Pi\<^isub>M UNIV M)" proof - have "Pi UNIV E = (\i. emb UNIV {..i} (\\<^isub>E j\{..i}. E j))" - using E E[THEN M.sets_into_space] - by (auto simp: emb_def Pi_iff extensional_def) blast - with E show ?thesis - by (auto intro: emb_in_infprod_algebra) + using E E[THEN sets_into_space] + by (auto simp: prod_emb_def Pi_iff extensional_def) blast + with E show ?thesis by auto qed -lemma (in sequence_space) measure_infprod: +lemma (in sequence_space) measure_PiM_countable: fixes E :: "nat \ 'a set" assumes E: "\i. E i \ sets (M i)" - shows "(\n. \i\n. M.\' i (E i)) ----> \' (Pi UNIV E)" + shows "(\n. \i\n. measure (M i) (E i)) ----> measure (Pi\<^isub>M UNIV M) (Pi UNIV E)" proof - let ?E = "\n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)" - { fix n :: nat - interpret n: finite_product_prob_space M "{..n}" by default auto - have "(\i\n. M.\' i (E i)) = n.\' (Pi\<^isub>E {.. n} E)" - using E by (subst n.finite_measure_times) auto - also have "\ = \' (?E n)" - using E by (intro finite_measure_infprod_emb[symmetric]) auto - finally have "(\i\n. M.\' i (E i)) = \' (?E n)" . } + have "\n. (\i\n. measure (M i) (E i)) = measure (Pi\<^isub>M UNIV M) (?E n)" + using E by (simp add: measure_PiM_emb) moreover have "Pi UNIV E = (\n. ?E n)" - using E E[THEN M.sets_into_space] - by (auto simp: emb_def extensional_def Pi_iff) blast - moreover have "range ?E \ sets (Pi\<^isub>P UNIV M)" + using E E[THEN sets_into_space] + by (auto simp: prod_emb_def extensional_def Pi_iff) blast + moreover have "range ?E \ sets (Pi\<^isub>M UNIV M)" using E by auto moreover have "decseq ?E" - by (auto simp: emb_def Pi_iff decseq_def) + by (auto simp: prod_emb_def Pi_iff decseq_def) ultimately show ?thesis - by (simp add: finite_continuity_from_above) + by (simp add: finite_Lim_measure_decseq) qed end \ No newline at end of file diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Information.thy Mon Apr 23 12:14:35 2012 +0200 @@ -172,147 +172,152 @@ Kullback$-$Leibler distance. *} definition - "entropy_density b M \ = log b \ real \ RN_deriv M \" + "entropy_density b M N = log b \ real \ RN_deriv M N" definition - "KL_divergence b M \ = integral\<^isup>L (M\measure := \\) (entropy_density b M \)" + "KL_divergence b M N = integral\<^isup>L N (entropy_density b M N)" lemma (in information_space) measurable_entropy_density: - assumes ps: "prob_space (M\measure := \\)" - assumes ac: "absolutely_continuous \" - shows "entropy_density b M \ \ borel_measurable M" + assumes ac: "absolutely_continuous M N" "sets N = events" + shows "entropy_density b M N \ borel_measurable M" proof - - interpret \: prob_space "M\measure := \\" by fact - have "measure_space (M\measure := \\)" by fact - from RN_deriv[OF this ac] b_gt_1 show ?thesis + from borel_measurable_RN_deriv[OF ac] b_gt_1 show ?thesis unfolding entropy_density_def by (intro measurable_comp) auto qed -lemma (in information_space) KL_gt_0: - assumes ps: "prob_space (M\measure := \\)" - assumes ac: "absolutely_continuous \" - assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" - assumes A: "A \ sets M" "\ A \ \ A" - shows "0 < KL_divergence b M \" -proof - - interpret \: prob_space "M\measure := \\" by fact - have ms: "measure_space (M\measure := \\)" by default - have fms: "finite_measure (M\measure := \\)" by unfold_locales - note RN = RN_deriv[OF ms ac] - - from real_RN_deriv[OF fms ac] guess D . note D = this - with absolutely_continuous_AE[OF ms] ac - have D\: "AE x in M\measure := \\. RN_deriv M \ x = ereal (D x)" - by auto - - def f \ "\x. if D x = 0 then 1 else 1 / D x" - with D have f_borel: "f \ borel_measurable M" - by (auto intro!: measurable_If) +lemma (in sigma_finite_measure) KL_density: + fixes f :: "'a \ real" + assumes "1 < b" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + shows "KL_divergence b M (density M f) = (\x. f x * log b (f x) \M)" + unfolding KL_divergence_def +proof (subst integral_density) + show "entropy_density b M (density M (\x. ereal (f x))) \ borel_measurable M" + using f `1 < b` + by (auto simp: comp_def entropy_density_def intro!: borel_measurable_log borel_measurable_RN_deriv_density) + have "density M (RN_deriv M (density M f)) = density M f" + using f by (intro density_RN_deriv_density) auto + then have eq: "AE x in M. RN_deriv M (density M f) x = f x" + using f + by (intro density_unique) + (auto intro!: borel_measurable_log borel_measurable_RN_deriv_density simp: RN_deriv_density_nonneg) + show "(\x. f x * entropy_density b M (density M (\x. ereal (f x))) x \M) = (\x. f x * log b (f x) \M)" + apply (intro integral_cong_AE) + using eq + apply eventually_elim + apply (auto simp: entropy_density_def) + done +qed fact+ - have "KL_divergence b M \ = 1 / ln b * (\ x. ln b * entropy_density b M \ x \M\measure := \\)" - unfolding KL_divergence_def using int b_gt_1 - by (simp add: integral_cmult) - - { fix A assume "A \ sets M" - with RN D have "\.\ A = (\\<^isup>+ x. ereal (D x) * indicator A x \M)" - by (auto intro!: positive_integral_cong_AE) } - note D_density = this +lemma (in sigma_finite_measure) KL_density_density: + fixes f g :: "'a \ real" + assumes "1 < b" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + assumes ac: "AE x in M. f x = 0 \ g x = 0" + shows "KL_divergence b (density M f) (density M g) = (\x. g x * log b (g x / f x) \M)" +proof - + interpret Mf: sigma_finite_measure "density M f" + using f by (subst sigma_finite_iff_density_finite) auto + have "KL_divergence b (density M f) (density M g) = + KL_divergence b (density M f) (density (density M f) (\x. g x / f x))" + using f g ac by (subst density_density_divide) simp_all + also have "\ = (\x. (g x / f x) * log b (g x / f x) \density M f)" + using f g `1 < b` by (intro Mf.KL_density) (auto simp: AE_density divide_nonneg_nonneg) + also have "\ = (\x. g x * log b (g x / f x) \M)" + using ac f g `1 < b` by (subst integral_density) (auto intro!: integral_cong_AE) + finally show ?thesis . +qed - have ln_entropy: "(\x. ln b * entropy_density b M \ x) \ borel_measurable M" - using measurable_entropy_density[OF ps ac] by auto +lemma (in information_space) KL_gt_0: + fixes D :: "'a \ real" + assumes "prob_space (density M D)" + assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" + assumes int: "integrable M (\x. D x * log b (D x))" + assumes A: "density M D \ M" + shows "0 < KL_divergence b M (density M D)" +proof - + interpret N: prob_space "density M D" by fact - have "integrable (M\measure := \\) (\x. ln b * entropy_density b M \ x)" - using int by auto - moreover have "integrable (M\measure := \\) (\x. ln b * entropy_density b M \ x) \ - integrable M (\x. D x * (ln b * entropy_density b M \ x))" - using D D_density ln_entropy - by (intro integral_translated_density) auto - ultimately have M_int: "integrable M (\x. D x * (ln b * entropy_density b M \ x))" - by simp + obtain A where "A \ sets M" "emeasure (density M D) A \ emeasure M A" + using measure_eqI[of "density M D" M] `density M D \ M` by auto + + let ?D_set = "{x\space M. D x \ 0}" + have [simp, intro]: "?D_set \ sets M" + using D by auto have D_neg: "(\\<^isup>+ x. ereal (- D x) \M) = 0" using D by (subst positive_integral_0_iff_AE) auto - have "(\\<^isup>+ x. ereal (D x) \M) = \ (space M)" - using RN D by (auto intro!: positive_integral_cong_AE) + have "(\\<^isup>+ x. ereal (D x) \M) = emeasure (density M D) (space M)" + using D by (simp add: emeasure_density cong: positive_integral_cong) then have D_pos: "(\\<^isup>+ x. ereal (D x) \M) = 1" - using \.measure_space_1 by simp - - have "integrable M D" - using D_pos D_neg D by (auto simp: integrable_def) + using N.emeasure_space_1 by simp - have "integral\<^isup>L M D = 1" - using D_pos D_neg by (auto simp: lebesgue_integral_def) + have "integrable M D" "integral\<^isup>L M D = 1" + using D D_pos D_neg unfolding integrable_def lebesgue_integral_def by simp_all - let ?D_set = "{x\space M. D x \ 0}" - have [simp, intro]: "?D_set \ sets M" - using D by (auto intro: sets_Collect) - - have "0 \ 1 - \' ?D_set" + have "0 \ 1 - measure M ?D_set" using prob_le_1 by (auto simp: field_simps) also have "\ = (\ x. D x - indicator ?D_set x \M)" using `integrable M D` `integral\<^isup>L M D = 1` - by (simp add: \'_def) - also have "\ < (\ x. D x * (ln b * entropy_density b M \ x) \M)" + by (simp add: emeasure_eq_measure) + also have "\ < (\ x. D x * (ln b * log b (D x)) \M)" proof (rule integral_less_AE) show "integrable M (\x. D x - indicator ?D_set x)" using `integrable M D` by (intro integral_diff integral_indicator) auto next - show "integrable M (\x. D x * (ln b * entropy_density b M \ x))" - by fact + from integral_cmult(1)[OF int, of "ln b"] + show "integrable M (\x. D x * (ln b * log b (D x)))" + by (simp add: ac_simps) next - show "\ {x\space M. D x \ 1 \ D x \ 0} \ 0" + show "emeasure M {x\space M. D x \ 1 \ D x \ 0} \ 0" proof - assume eq_0: "\ {x\space M. D x \ 1 \ D x \ 0} = 0" - then have disj: "AE x. D x = 1 \ D x = 0" + assume eq_0: "emeasure M {x\space M. D x \ 1 \ D x \ 0} = 0" + then have disj: "AE x in M. D x = 1 \ D x = 0" using D(1) by (auto intro!: AE_I[OF subset_refl] sets_Collect) - have "\ {x\space M. D x = 1} = (\\<^isup>+ x. indicator {x\space M. D x = 1} x \M)" + have "emeasure M {x\space M. D x = 1} = (\\<^isup>+ x. indicator {x\space M. D x = 1} x \M)" using D(1) by auto - also have "\ = (\\<^isup>+ x. ereal (D x) * indicator {x\space M. D x \ 0} x \M)" + also have "\ = (\\<^isup>+ x. ereal (D x) \M)" using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def) - also have "\ = \ {x\space M. D x \ 0}" - using D(1) D_density by auto - also have "\ = \ (space M)" - using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def) - finally have "AE x. D x = 1" - using D(1) \.measure_space_1 by (intro AE_I_eq_1) auto + finally have "AE x in M. D x = 1" + using D D_pos by (intro AE_I_eq_1) auto then have "(\\<^isup>+x. indicator A x\M) = (\\<^isup>+x. ereal (D x) * indicator A x\M)" by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric]) - also have "\ = \ A" - using `A \ sets M` D_density by simp - finally show False using `A \ sets M` `\ A \ \ A` by simp + also have "\ = density M D A" + using `A \ sets M` D by (simp add: emeasure_density) + finally show False using `A \ sets M` `emeasure (density M D) A \ emeasure M A` by simp qed show "{x\space M. D x \ 1 \ D x \ 0} \ sets M" - using D(1) by (auto intro: sets_Collect) + using D(1) by (auto intro: sets_Collect_conj) - show "AE t. t \ {x\space M. D x \ 1 \ D x \ 0} \ - D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" + show "AE t in M. t \ {x\space M. D x \ 1 \ D x \ 0} \ + D t - indicator ?D_set t \ D t * (ln b * log b (D t))" using D(2) - proof (elim AE_mp, safe intro!: AE_I2) - fix t assume Dt: "t \ space M" "D t \ 1" "D t \ 0" - and RN: "RN_deriv M \ t = ereal (D t)" - and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \ t)" + proof (eventually_elim, safe) + fix t assume Dt: "t \ space M" "D t \ 1" "D t \ 0" "0 \ D t" + and eq: "D t - indicator ?D_set t = D t * (ln b * log b (D t))" have "D t - 1 = D t - indicator ?D_set t" using Dt by simp also note eq - also have "D t * (ln b * entropy_density b M \ t) = - D t * ln (1 / D t)" - using RN b_gt_1 `D t \ 0` `0 \ D t` - by (simp add: entropy_density_def log_def ln_div less_le) + also have "D t * (ln b * log b (D t)) = - D t * ln (1 / D t)" + using b_gt_1 `D t \ 0` `0 \ D t` + by (simp add: log_def ln_div less_le) finally have "ln (1 / D t) = 1 / D t - 1" using `D t \ 0` by (auto simp: field_simps) from ln_eq_minus_one[OF _ this] `D t \ 0` `0 \ D t` `D t \ 1` show False by auto qed - show "AE t. D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" - using D(2) - proof (elim AE_mp, intro AE_I2 impI) - fix t assume "t \ space M" and RN: "RN_deriv M \ t = ereal (D t)" - show "D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" + show "AE t in M. D t - indicator ?D_set t \ D t * (ln b * log b (D t))" + using D(2) AE_space + proof eventually_elim + fix t assume "t \ space M" "0 \ D t" + show "D t - indicator ?D_set t \ D t * (ln b * log b (D t))" proof cases assume asm: "D t \ 0" then have "0 < D t" using `0 \ D t` by auto @@ -321,592 +326,425 @@ using asm `t \ space M` by (simp add: field_simps) also have "- D t * (1 / D t - 1) \ - D t * ln (1 / D t)" using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto - also have "\ = D t * (ln b * entropy_density b M \ t)" - using `0 < D t` RN b_gt_1 - by (simp_all add: log_def ln_div entropy_density_def) + also have "\ = D t * (ln b * log b (D t))" + using `0 < D t` b_gt_1 + by (simp_all add: log_def ln_div) finally show ?thesis by simp qed simp qed qed - also have "\ = (\ x. ln b * entropy_density b M \ x \M\measure := \\)" - using D D_density ln_entropy - by (intro integral_translated_density[symmetric]) auto - also have "\ = ln b * (\ x. entropy_density b M \ x \M\measure := \\)" - using int by (rule \.integral_cmult) - finally show "0 < KL_divergence b M \" - using b_gt_1 by (auto simp: KL_divergence_def zero_less_mult_iff) + also have "\ = (\ x. ln b * (D x * log b (D x)) \M)" + by (simp add: ac_simps) + also have "\ = ln b * (\ x. D x * log b (D x) \M)" + using int by (rule integral_cmult) + finally show ?thesis + using b_gt_1 D by (subst KL_density) (auto simp: zero_less_mult_iff) qed -lemma (in sigma_finite_measure) KL_eq_0: - assumes eq: "\A\sets M. \ A = measure M A" - shows "KL_divergence b M \ = 0" +lemma (in sigma_finite_measure) KL_same_eq_0: "KL_divergence b M M = 0" proof - - have "AE x. 1 = RN_deriv M \ x" + have "AE x in M. 1 = RN_deriv M M x" proof (rule RN_deriv_unique) - show "measure_space (M\measure := \\)" - using eq by (intro measure_space_cong) auto - show "absolutely_continuous \" - unfolding absolutely_continuous_def using eq by auto - show "(\x. 1) \ borel_measurable M" "AE x. 0 \ (1 :: ereal)" by auto - fix A assume "A \ sets M" - with eq show "\ A = \\<^isup>+ x. 1 * indicator A x \M" by simp + show "(\x. 1) \ borel_measurable M" "AE x in M. 0 \ (1 :: ereal)" by auto + show "density M (\x. 1) = M" + apply (auto intro!: measure_eqI emeasure_density) + apply (subst emeasure_density) + apply auto + done qed - then have "AE x. log b (real (RN_deriv M \ x)) = 0" + then have "AE x in M. log b (real (RN_deriv M M x)) = 0" by (elim AE_mp) simp from integral_cong_AE[OF this] - have "integral\<^isup>L M (entropy_density b M \) = 0" + have "integral\<^isup>L M (entropy_density b M M) = 0" by (simp add: entropy_density_def comp_def) - with eq show "KL_divergence b M \ = 0" + then show "KL_divergence b M M = 0" unfolding KL_divergence_def - by (subst integral_cong_measure) auto + by auto qed -lemma (in information_space) KL_eq_0_imp: - assumes ps: "prob_space (M\measure := \\)" - assumes ac: "absolutely_continuous \" - assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" - assumes KL: "KL_divergence b M \ = 0" - shows "\A\sets M. \ A = \ A" - by (metis less_imp_neq KL_gt_0 assms) +lemma (in information_space) KL_eq_0_iff_eq: + fixes D :: "'a \ real" + assumes "prob_space (density M D)" + assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" + assumes int: "integrable M (\x. D x * log b (D x))" + shows "KL_divergence b M (density M D) = 0 \ density M D = M" + using KL_same_eq_0[of b] KL_gt_0[OF assms] + by (auto simp: less_le) -lemma (in information_space) KL_ge_0: - assumes ps: "prob_space (M\measure := \\)" - assumes ac: "absolutely_continuous \" - assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" - shows "0 \ KL_divergence b M \" - using KL_eq_0 KL_gt_0[OF ps ac int] - by (cases "\A\sets M. \ A = measure M A") (auto simp: le_less) - - -lemma (in sigma_finite_measure) KL_divergence_vimage: - assumes T: "T \ measure_preserving M M'" - and T': "T' \ measure_preserving (M'\ measure := \' \) (M\ measure := \ \)" - and inv: "\x. x \ space M \ T' (T x) = x" - and inv': "\x. x \ space M' \ T (T' x) = x" - and \': "measure_space (M'\measure := \'\)" "measure_space.absolutely_continuous M' \'" - and "1 < b" - shows "KL_divergence b M' \' = KL_divergence b M \" +lemma (in information_space) KL_eq_0_iff_eq_ac: + fixes D :: "'a \ real" + assumes "prob_space N" + assumes ac: "absolutely_continuous M N" "sets N = sets M" + assumes int: "integrable N (entropy_density b M N)" + shows "KL_divergence b M N = 0 \ N = M" proof - - interpret \': measure_space "M'\measure := \'\" by fact - have M: "measure_space (M\ measure := \\)" - by (rule \'.measure_space_vimage[OF _ T'], simp) default - have "sigma_algebra (M'\ measure := \'\)" by default - then have saM': "sigma_algebra M'" by simp - then interpret M': measure_space M' by (rule measure_space_vimage) fact - have ac: "absolutely_continuous \" unfolding absolutely_continuous_def - proof safe - fix N assume N: "N \ sets M" and N_0: "\ N = 0" - then have N': "T' -` N \ space M' \ sets M'" - using T' by (auto simp: measurable_def measure_preserving_def) - have "T -` (T' -` N \ space M') \ space M = N" - using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def) - then have "measure M' (T' -` N \ space M') = 0" - using measure_preservingD[OF T N'] N_0 by auto - with \'(2) N' show "\ N = 0" using measure_preservingD[OF T', of N] N - unfolding M'.absolutely_continuous_def measurable_def by auto - qed + interpret N: prob_space N by fact + have "finite_measure N" by unfold_locales + from real_RN_deriv[OF this ac] guess D . note D = this + + have "N = density M (RN_deriv M N)" + using ac by (rule density_RN_deriv[symmetric]) + also have "\ = density M D" + using borel_measurable_RN_deriv[OF ac] D by (auto intro!: density_cong) + finally have N: "N = density M D" . - have sa: "sigma_algebra (M\measure := \\)" by simp default - have AE: "AE x. RN_deriv M' \' (T x) = RN_deriv M \ x" - by (rule RN_deriv_vimage[OF T T' inv \']) - show ?thesis - unfolding KL_divergence_def entropy_density_def comp_def - proof (subst \'.integral_vimage[OF sa T']) - show "(\x. log b (real (RN_deriv M \ x))) \ borel_measurable (M\measure := \\)" - by (auto intro!: RN_deriv[OF M ac] borel_measurable_log[OF _ `1 < b`]) - have "(\ x. log b (real (RN_deriv M' \' x)) \M'\measure := \'\) = - (\ x. log b (real (RN_deriv M' \' (T (T' x)))) \M'\measure := \'\)" (is "?l = _") - using inv' by (auto intro!: \'.integral_cong) - also have "\ = (\ x. log b (real (RN_deriv M \ (T' x))) \M'\measure := \'\)" (is "_ = ?r") - using M ac AE - by (intro \'.integral_cong_AE \'.almost_everywhere_vimage[OF sa T'] absolutely_continuous_AE[OF M]) - (auto elim!: AE_mp) - finally show "?l = ?r" . - qed + from absolutely_continuous_AE[OF ac(2,1) D(2)] D b_gt_1 ac measurable_entropy_density + have "integrable N (\x. log b (D x))" + by (intro integrable_cong_AE[THEN iffD2, OF _ _ _ int]) + (auto simp: N entropy_density_def) + with D b_gt_1 have "integrable M (\x. D x * log b (D x))" + by (subst integral_density(2)[symmetric]) (auto simp: N[symmetric] comp_def) + with `prob_space N` D show ?thesis + unfolding N + by (intro KL_eq_0_iff_eq) auto qed -lemma (in sigma_finite_measure) KL_divergence_cong: - 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 ?\ 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 entropy_density_def) - also have "\ = KL_divergence b N \'" - by (auto intro!: \.integral_cong_measure[symmetric] simp: KL_divergence_def entropy_density_def comp_def) - finally show ?thesis . -qed +lemma (in information_space) KL_nonneg: + assumes "prob_space (density M D)" + assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" + assumes int: "integrable M (\x. D x * log b (D x))" + shows "0 \ KL_divergence b M (density M D)" + using KL_gt_0[OF assms] by (cases "density M D = M") (auto simp: KL_same_eq_0) -lemma (in finite_measure_space) KL_divergence_eq_finite: - 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") -proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v] entropy_density_def) - 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) -qed +lemma (in sigma_finite_measure) KL_density_density_nonneg: + fixes f g :: "'a \ real" + assumes "1 < b" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" "prob_space (density M f)" + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" "prob_space (density M g)" + assumes ac: "AE x in M. f x = 0 \ g x = 0" + assumes int: "integrable M (\x. g x * log b (g x / f x))" + shows "0 \ KL_divergence b (density M f) (density M g)" +proof - + interpret Mf: prob_space "density M f" by fact + interpret Mf: information_space "density M f" b by default fact + have eq: "density (density M f) (\x. g x / f x) = density M g" (is "?DD = _") + using f g ac by (subst density_density_divide) simp_all -lemma (in finite_prob_space) KL_divergence_positive_finite: - assumes v: "finite_prob_space (M\measure := \\)" - assumes ac: "absolutely_continuous \" - and "1 < b" - shows "0 \ KL_divergence b M \" -proof - - interpret information_space M by default fact - interpret v: finite_prob_space "M\measure := \\" by fact - have ps: "prob_space (M\measure := \\)" by unfold_locales - from KL_ge_0[OF this ac v.integral_finite_singleton(1)] show ?thesis . + have "0 \ KL_divergence b (density M f) (density (density M f) (\x. g x / f x))" + proof (rule Mf.KL_nonneg) + show "prob_space ?DD" unfolding eq by fact + from f g show "(\x. g x / f x) \ borel_measurable (density M f)" + by auto + show "AE x in density M f. 0 \ g x / f x" + using f g by (auto simp: AE_density divide_nonneg_nonneg) + show "integrable (density M f) (\x. g x / f x * log b (g x / f x))" + using `1 < b` f g ac + by (subst integral_density) + (auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If) + qed + also have "\ = KL_divergence b (density M f) (density M g)" + using f g ac by (subst density_density_divide) simp_all + finally show ?thesis . qed subsection {* Mutual Information *} definition (in prob_space) "mutual_information b S T X Y = - KL_divergence b (S\measure := ereal\distribution X\ \\<^isub>M T\measure := ereal\distribution Y\) - (ereal\joint_distribution X Y)" + KL_divergence b (distr M S X \\<^isub>M distr M T Y) (distr M (S \\<^isub>M T) (\x. (X x, Y x)))" -lemma (in information_space) +lemma (in information_space) mutual_information_indep_vars: fixes S T X Y - defines "P \ S\measure := ereal\distribution X\ \\<^isub>M T\measure := ereal\distribution Y\" + defines "P \ distr M S X \\<^isub>M distr M T Y" + defines "Q \ distr M (S \\<^isub>M T) (\x. (X x, Y x))" shows "indep_var S X T Y \ (random_variable S X \ random_variable T Y \ - measure_space.absolutely_continuous P (ereal\joint_distribution X Y) \ - integrable (P\measure := (ereal\joint_distribution X Y)\) - (entropy_density b P (ereal\joint_distribution X Y)) \ - mutual_information b S T X Y = 0)" + absolutely_continuous P Q \ integrable Q (entropy_density b P Q) \ + mutual_information b S T X Y = 0)" + unfolding indep_var_distribution_eq proof safe - assume indep: "indep_var S X T Y" - then have "random_variable S X" "random_variable T Y" - by (blast dest: indep_var_rv1 indep_var_rv2)+ - then show "sigma_algebra S" "X \ measurable M S" "sigma_algebra T" "Y \ measurable M T" - by blast+ - - interpret X: prob_space "S\measure := ereal\distribution X\" - by (rule distribution_prob_space) fact - interpret Y: prob_space "T\measure := ereal\distribution Y\" - by (rule distribution_prob_space) fact - interpret XY: pair_prob_space "S\measure := ereal\distribution X\" "T\measure := ereal\distribution Y\" by default - interpret XY: information_space XY.P b by default (rule b_gt_1) - - let ?J = "XY.P\ measure := (ereal\joint_distribution X Y) \" - { fix A assume "A \ sets XY.P" - then have "ereal (joint_distribution X Y A) = XY.\ A" - using indep_var_distributionD[OF indep] - by (simp add: XY.P.finite_measure_eq) } - note j_eq = this - - interpret J: prob_space ?J - using j_eq by (intro XY.prob_space_cong) auto - - have ac: "XY.absolutely_continuous (ereal\joint_distribution X Y)" - by (simp add: XY.absolutely_continuous_def j_eq) - then show "measure_space.absolutely_continuous P (ereal\joint_distribution X Y)" - unfolding P_def . - - have ed: "entropy_density b XY.P (ereal\joint_distribution X Y) \ borel_measurable XY.P" - by (rule XY.measurable_entropy_density) (default | fact)+ + assume rv: "random_variable S X" "random_variable T Y" - have "AE x in XY.P. 1 = RN_deriv XY.P (ereal\joint_distribution X Y) x" - proof (rule XY.RN_deriv_unique[OF _ ac]) - show "measure_space ?J" by default - fix A assume "A \ sets XY.P" - then show "(ereal\joint_distribution X Y) A = (\\<^isup>+ x. 1 * indicator A x \XY.P)" - by (simp add: j_eq) - qed (insert XY.measurable_const[of 1 borel], auto) - then have ae_XY: "AE x in XY.P. entropy_density b XY.P (ereal\joint_distribution X Y) x = 0" - by (elim XY.AE_mp) (simp add: entropy_density_def) - have ae_J: "AE x in ?J. entropy_density b XY.P (ereal\joint_distribution X Y) x = 0" - proof (rule XY.absolutely_continuous_AE) - show "measure_space ?J" by default - show "XY.absolutely_continuous (measure ?J)" - using ac by simp - qed (insert ae_XY, simp_all) - then show "integrable (P\measure := (ereal\joint_distribution X Y)\) - (entropy_density b P (ereal\joint_distribution X Y))" - unfolding P_def - using ed XY.measurable_const[of 0 borel] - by (subst J.integrable_cong_AE) auto + interpret X: prob_space "distr M S X" + by (rule prob_space_distr) fact + interpret Y: prob_space "distr M T Y" + by (rule prob_space_distr) fact + interpret XY: pair_prob_space "distr M S X" "distr M T Y" by default + interpret P: information_space P b unfolding P_def by default (rule b_gt_1) - show "mutual_information b S T X Y = 0" - unfolding mutual_information_def KL_divergence_def P_def - by (subst J.integral_cong_AE[OF ae_J]) simp -next - assume "sigma_algebra S" "X \ measurable M S" "sigma_algebra T" "Y \ measurable M T" - then have rvs: "random_variable S X" "random_variable T Y" by blast+ + interpret Q: prob_space Q unfolding Q_def + by (rule prob_space_distr) (simp add: comp_def measurable_pair_iff rv) - interpret X: prob_space "S\measure := ereal\distribution X\" - by (rule distribution_prob_space) fact - interpret Y: prob_space "T\measure := ereal\distribution Y\" - by (rule distribution_prob_space) fact - interpret XY: pair_prob_space "S\measure := ereal\distribution X\" "T\measure := ereal\distribution Y\" by default - interpret XY: information_space XY.P b by default (rule b_gt_1) + { assume "distr M S X \\<^isub>M distr M T Y = distr M (S \\<^isub>M T) (\x. (X x, Y x))" + then have [simp]: "Q = P" unfolding Q_def P_def by simp - let ?J = "XY.P\ measure := (ereal\joint_distribution X Y) \" - interpret J: prob_space ?J - using rvs by (intro joint_distribution_prob_space) auto + show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def) + then have ed: "entropy_density b P Q \ borel_measurable P" + by (rule P.measurable_entropy_density) simp - assume ac: "measure_space.absolutely_continuous P (ereal\joint_distribution X Y)" - assume int: "integrable (P\measure := (ereal\joint_distribution X Y)\) - (entropy_density b P (ereal\joint_distribution X Y))" - assume I_eq_0: "mutual_information b S T X Y = 0" + have "AE x in P. 1 = RN_deriv P Q x" + proof (rule P.RN_deriv_unique) + show "density P (\x. 1) = Q" + unfolding `Q = P` by (intro measure_eqI) (auto simp: emeasure_density) + qed auto + then have ae_0: "AE x in P. entropy_density b P Q x = 0" + by eventually_elim (auto simp: entropy_density_def) + then have "integrable P (entropy_density b P Q) \ integrable Q (\x. 0)" + using ed unfolding `Q = P` by (intro integrable_cong_AE) auto + then show "integrable Q (entropy_density b P Q)" by simp - have eq: "\A\sets XY.P. (ereal \ joint_distribution X Y) A = XY.\ A" - proof (rule XY.KL_eq_0_imp) - show "prob_space ?J" by unfold_locales - show "XY.absolutely_continuous (ereal\joint_distribution X Y)" - using ac by (simp add: P_def) - show "integrable ?J (entropy_density b XY.P (ereal\joint_distribution X Y))" - using int by (simp add: P_def) - show "KL_divergence b XY.P (ereal\joint_distribution X Y) = 0" - using I_eq_0 unfolding mutual_information_def by (simp add: P_def) - qed - - { fix S X assume "sigma_algebra S" - interpret S: sigma_algebra S by fact - have "Int_stable \space = space M, sets = {X -` A \ space M |A. A \ sets S}\" - proof (safe intro!: Int_stableI) - fix A B assume "A \ sets S" "B \ sets S" - then show "\C. (X -` A \ space M) \ (X -` B \ space M) = (X -` C \ space M) \ C \ sets S" - by (intro exI[of _ "A \ B"]) auto - qed } - note Int_stable = this + show "mutual_information b S T X Y = 0" + unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] `Q = P` + using ae_0 by (simp cong: integral_cong_AE) } - show "indep_var S X T Y" unfolding indep_var_eq - proof (intro conjI indep_set_sigma_sets Int_stable) - show "indep_set {X -` A \ space M |A. A \ sets S} {Y -` A \ space M |A. A \ sets T}" - proof (safe intro!: indep_setI) - { fix A assume "A \ sets S" then show "X -` A \ space M \ sets M" - using `X \ measurable M S` by (auto intro: measurable_sets) } - { fix A assume "A \ sets T" then show "Y -` A \ space M \ sets M" - using `Y \ measurable M T` by (auto intro: measurable_sets) } - next - fix A B assume ab: "A \ sets S" "B \ sets T" - have "ereal (prob ((X -` A \ space M) \ (Y -` B \ space M))) = - ereal (joint_distribution X Y (A \ B))" - unfolding distribution_def - by (intro arg_cong[where f="\C. ereal (prob C)"]) auto - also have "\ = XY.\ (A \ B)" - using ab eq by (auto simp: XY.finite_measure_eq) - also have "\ = ereal (distribution X A) * ereal (distribution Y B)" - using ab by (simp add: XY.pair_measure_times) - finally show "prob ((X -` A \ space M) \ (Y -` B \ space M)) = - prob (X -` A \ space M) * prob (Y -` B \ space M)" - unfolding distribution_def by simp - qed - qed fact+ -qed + { assume ac: "absolutely_continuous P Q" + assume int: "integrable Q (entropy_density b P Q)" + assume I_eq_0: "mutual_information b S T X Y = 0" -lemma (in information_space) mutual_information_commute_generic: - assumes X: "random_variable S X" and Y: "random_variable T Y" - assumes ac: "measure_space.absolutely_continuous - (S\measure := ereal\distribution X\ \\<^isub>M T\measure := ereal\distribution Y\) (ereal\joint_distribution X Y)" - shows "mutual_information b S T X Y = mutual_information b T S Y X" -proof - - let ?S = "S\measure := ereal\distribution X\" and ?T = "T\measure := ereal\distribution Y\" - interpret S: prob_space ?S using X by (rule distribution_prob_space) - interpret T: prob_space ?T using Y by (rule distribution_prob_space) - interpret P: pair_prob_space ?S ?T .. - interpret Q: pair_prob_space ?T ?S .. - show ?thesis - unfolding mutual_information_def - proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1]) - show "(\(x,y). (y,x)) \ measure_preserving - (P.P \ measure := ereal\joint_distribution X Y\) (Q.P \ measure := ereal\joint_distribution Y X\)" - using X Y unfolding measurable_def - unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable - by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\']) - have "prob_space (P.P\ measure := ereal\joint_distribution X Y\)" - using X Y by (auto intro!: distribution_prob_space random_variable_pairI) - then show "measure_space (P.P\ measure := ereal\joint_distribution X Y\)" - unfolding prob_space_def finite_measure_def sigma_finite_measure_def by simp - qed auto + have eq: "Q = P" + proof (rule P.KL_eq_0_iff_eq_ac[THEN iffD1]) + show "prob_space Q" by unfold_locales + show "absolutely_continuous P Q" by fact + show "integrable Q (entropy_density b P Q)" by fact + show "sets Q = sets P" by (simp add: P_def Q_def sets_pair_measure) + show "KL_divergence b P Q = 0" + using I_eq_0 unfolding mutual_information_def by (simp add: P_def Q_def) + qed + then show "distr M S X \\<^isub>M distr M T Y = distr M (S \\<^isub>M T) (\x. (X x, Y x))" + unfolding P_def Q_def .. } qed -definition (in prob_space) - "entropy b s X = mutual_information b s s X X" - abbreviation (in information_space) mutual_information_Pow ("\'(_ ; _')") where - "\(X ; Y) \ mutual_information b - \ space = X`space M, sets = Pow (X`space M), measure = ereal\distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = ereal\distribution Y \ X Y" + "\(X ; Y) \ mutual_information b (count_space (X`space M)) (count_space (Y`space M)) X Y" -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 - (S\measure := ereal\distribution X\ \\<^isub>M T\measure := ereal\distribution Y\) - (ereal\joint_distribution X Y)" +lemma (in information_space) + fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + defines "f \ \x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" + shows mutual_information_distr: "mutual_information b S T X Y = integral\<^isup>L (S \\<^isub>M T) f" (is "?M = ?R") + and mutual_information_nonneg: "integrable (S \\<^isub>M T) f \ 0 \ mutual_information b S T X Y" proof - - interpret X: finite_prob_space "S\measure := ereal\distribution X\" - using X by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space "T\measure := ereal\distribution Y\" - using Y by (rule distribution_finite_prob_space) - interpret XY: pair_finite_prob_space - "S\measure := ereal\distribution X\" "T\ measure := ereal\distribution Y\" by default - interpret P: finite_prob_space "XY.P\ measure := ereal\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 (ereal\joint_distribution X Y)" - proof (rule XY.absolutely_continuousI) - show "finite_measure_space (XY.P\ measure := ereal\joint_distribution X Y\)" by unfold_locales - fix x assume "x \ space XY.P" and "XY.\ {x} = 0" - then obtain a b where "x = (a, b)" - and "distribution X {a} = 0 \ distribution Y {b} = 0" - by (cases x) (auto simp: space_pair_measure) - with finite_distribution_order(5,6)[OF X Y] - show "(ereal \ joint_distribution X Y) {x} = 0" by auto + have X: "random_variable S X" + using Px by (auto simp: distributed_def) + have Y: "random_variable T Y" + using Py by (auto simp: distributed_def) + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T .. + interpret X: prob_space "distr M S X" using X by (rule prob_space_distr) + interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr) + interpret XY: pair_prob_space "distr M S X" "distr M T Y" .. + let ?P = "S \\<^isub>M T" + let ?D = "distr M ?P (\x. (X x, Y x))" + + { fix A assume "A \ sets S" + with X Y have "emeasure (distr M S X) A = emeasure ?D (A \ space T)" + by (auto simp: emeasure_distr measurable_Pair measurable_space + intro!: arg_cong[where f="emeasure M"]) } + note marginal_eq1 = this + { fix A assume "A \ sets T" + with X Y have "emeasure (distr M T Y) A = emeasure ?D (space S \ A)" + by (auto simp: emeasure_distr measurable_Pair measurable_space + intro!: arg_cong[where f="emeasure M"]) } + note marginal_eq2 = this + + have eq: "(\x. ereal (Px (fst x) * Py (snd x))) = (\(x, y). ereal (Px x) * ereal (Py y))" + by auto + + have distr_eq: "distr M S X \\<^isub>M distr M T Y = density ?P (\x. ereal (Px (fst x) * Py (snd x)))" + unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq + proof (subst pair_measure_density) + show "(\x. ereal (Px x)) \ borel_measurable S" "(\y. ereal (Py y)) \ borel_measurable T" + "AE x in S. 0 \ ereal (Px x)" "AE y in T. 0 \ ereal (Py y)" + using Px Py by (auto simp: distributed_def) + show "sigma_finite_measure (density S Px)" unfolding Px(1)[THEN distributed_distr_eq_density, symmetric] .. + show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] .. + qed (fact | simp)+ + + have M: "?M = KL_divergence b (density ?P (\x. ereal (Px (fst x) * Py (snd x)))) (density ?P (\x. ereal (Pxy x)))" + unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] .. + + from Px Py have f: "(\x. Px (fst x) * Py (snd x)) \ borel_measurable ?P" + by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'') + have PxPy_nonneg: "AE x in ?P. 0 \ Px (fst x) * Py (snd x)" + proof (rule ST.AE_pair_measure) + show "{x \ space ?P. 0 \ Px (fst x) * Py (snd x)} \ sets ?P" + using f by auto + show "AE x in S. AE y in T. 0 \ Px (fst (x, y)) * Py (snd (x, y))" + using Px Py by (auto simp: zero_le_mult_iff dest!: distributed_real_AE) + qed + + have "(AE x in ?P. Px (fst x) = 0 \ Pxy x = 0)" + by (rule subdensity_real[OF measurable_fst Pxy Px]) auto + moreover + have "(AE x in ?P. Py (snd x) = 0 \ Pxy x = 0)" + by (rule subdensity_real[OF measurable_snd Pxy Py]) auto + ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \ Pxy x = 0" + by eventually_elim auto + + show "?M = ?R" + unfolding M f_def + using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac + by (rule ST.KL_density_density) + + assume int: "integrable (S \\<^isub>M T) f" + show "0 \ ?M" unfolding M + proof (rule ST.KL_density_density_nonneg + [OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]]) + show "prob_space (density (S \\<^isub>M T) (\x. ereal (Pxy x))) " + unfolding distributed_distr_eq_density[OF Pxy, symmetric] + using distributed_measurable[OF Pxy] by (rule prob_space_distr) + show "prob_space (density (S \\<^isub>M T) (\x. ereal (Px (fst x) * Py (snd x))))" + unfolding distr_eq[symmetric] by unfold_locales qed qed lemma (in information_space) - assumes MX: "finite_random_variable MX X" - assumes MY: "finite_random_variable MY Y" - shows mutual_information_generic_eq: - "mutual_information b MX MY X Y = (\ (x,y) \ space MX \ space MY. - joint_distribution X Y {(x,y)} * - log b (joint_distribution X Y {(x,y)} / - (distribution X {x} * distribution Y {y})))" - (is ?sum) - and mutual_information_positive_generic: - "0 \ mutual_information b MX MY X Y" (is ?positive) + fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y" + shows mutual_information_eq_0: "mutual_information b S T X Y = 0" proof - - interpret X: finite_prob_space "MX\measure := ereal\distribution X\" - using MX by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space "MY\measure := ereal\distribution Y\" - using MY by (rule distribution_finite_prob_space) - interpret XY: pair_finite_prob_space "MX\measure := ereal\distribution X\" "MY\measure := ereal\distribution Y\" by default - interpret P: finite_prob_space "XY.P\measure := ereal\joint_distribution X Y\" - using assms by (auto intro!: joint_distribution_finite_prob_space) + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T .. - have P_ms: "finite_measure_space (XY.P\measure := ereal\joint_distribution X Y\)" by unfold_locales - have P_ps: "finite_prob_space (XY.P\measure := ereal\joint_distribution X Y\)" by unfold_locales - - 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: space_pair_measure setsum_cartesian_product') - - show ?positive - using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1] - unfolding mutual_information_def . + have "AE x in S \\<^isub>M T. Px (fst x) = 0 \ Pxy x = 0" + by (rule subdensity_real[OF measurable_fst Pxy Px]) auto + moreover + have "AE x in S \\<^isub>M T. Py (snd x) = 0 \ Pxy x = 0" + by (rule subdensity_real[OF measurable_snd Pxy Py]) auto + moreover + have "AE x in S \\<^isub>M T. Pxy x = Px (fst x) * Py (snd x)" + using distributed_real_measurable[OF Px] distributed_real_measurable[OF Py] distributed_real_measurable[OF Pxy] + by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'') + ultimately have "AE x in S \\<^isub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0" + by eventually_elim simp + then have "(\x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \(S \\<^isub>M T)) = (\x. 0 \(S \\<^isub>M T))" + by (rule integral_cong_AE) + then show ?thesis + by (subst mutual_information_distr[OF assms(1-5)]) simp qed -lemma (in information_space) mutual_information_commute: - assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" - shows "mutual_information b S T X Y = mutual_information b T S Y X" - unfolding mutual_information_generic_eq[OF X Y] mutual_information_generic_eq[OF Y X] - unfolding joint_distribution_commute_singleton[of X Y] - 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 M X" and Y: "simple_function M Y" - shows "\(X;Y) = \(Y;X)" - by (intro mutual_information_commute X Y simple_function_imp_finite_random_variable) - -lemma (in information_space) mutual_information_eq: - assumes "simple_function M X" "simple_function M Y" - shows "\(X;Y) = (\ (x,y) \ X ` space M \ Y ` space M. - distribution (\x. (X x, Y x)) {(x,y)} * log b (distribution (\x. (X x, Y x)) {(x,y)} / - (distribution X {x} * distribution Y {y})))" - using assms by (simp add: mutual_information_generic_eq) +lemma (in information_space) mutual_information_simple_distributed: + assumes X: "simple_distributed M X Px" and Y: "simple_distributed M Y Py" + assumes XY: "simple_distributed M (\x. (X x, Y x)) Pxy" + shows "\(X ; Y) = (\(x, y)\(\x. (X x, Y x))`space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" +proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]]) + note fin = simple_distributed_joint_finite[OF XY, simp] + show "sigma_finite_measure (count_space (X ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + show "sigma_finite_measure (count_space (Y ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + let ?Pxy = "\x. (if x \ (\x. (X x, Y x)) ` space M then Pxy x else 0)" + let ?f = "\x. ?Pxy x * log b (?Pxy x / (Px (fst x) * Py (snd x)))" + have "\x. ?f x = (if x \ (\x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) else 0)" + by auto + with fin show "(\ x. ?f x \(count_space (X ` space M) \\<^isub>M count_space (Y ` space M))) = + (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" + by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum_cases split_beta' + intro!: setsum_cong) +qed -lemma (in information_space) mutual_information_generic_cong: - assumes X: "\x. x \ space M \ X x = X' x" - assumes Y: "\x. x \ space M \ Y x = Y' x" - shows "mutual_information b MX MY X Y = mutual_information b MX MY X' Y'" - unfolding mutual_information_def using X Y - by (simp cong: distribution_cong) - -lemma (in information_space) mutual_information_cong: - assumes X: "\x. x \ space M \ X x = X' x" - assumes Y: "\x. x \ space M \ Y x = Y' x" - shows "\(X; Y) = \(X'; Y')" - unfolding mutual_information_def using X Y - by (simp cong: distribution_cong image_cong) - -lemma (in information_space) mutual_information_positive: - assumes "simple_function M X" "simple_function M Y" - shows "0 \ \(X;Y)" - using assms by (simp add: mutual_information_positive_generic) +lemma (in information_space) + fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" + assumes Px: "simple_distributed M X Px" and Py: "simple_distributed M Y Py" + assumes Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" + assumes ae: "\x\space M. Pxy (X x, Y x) = Px (X x) * Py (Y x)" + shows mutual_information_eq_0_simple: "\(X ; Y) = 0" +proof (subst mutual_information_simple_distributed[OF Px Py Pxy]) + have "(\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = + (\(x, y)\(\x. (X x, Y x)) ` space M. 0)" + by (intro setsum_cong) (auto simp: ae) + then show "(\(x, y)\(\x. (X x, Y x)) ` space M. + Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp +qed subsection {* Entropy *} +definition (in prob_space) entropy :: "real \ 'b measure \ ('a \ 'b) \ real" where + "entropy b S X = - KL_divergence b S (distr M S X)" + abbreviation (in information_space) entropy_Pow ("\'(_')") where - "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M), measure = ereal\distribution X \ X" - -lemma (in information_space) entropy_generic_eq: - fixes X :: "'a \ 'c" - assumes MX: "finite_random_variable MX X" - shows "entropy b MX X = -(\ x \ space MX. distribution X {x} * log b (distribution X {x}))" -proof - - interpret MX: finite_prob_space "MX\measure := ereal\distribution X\" - using MX by (rule distribution_finite_prob_space) - let ?X = "\x. distribution X {x}" - let ?XX = "\x y. joint_distribution X X {(x, y)}" + "\(X) \ entropy b (count_space (X`space M)) X" - { fix x y :: 'c - { assume "x \ y" - then have "(\x. (X x, X x)) -` {(x,y)} \ space M = {}" by auto - then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) } - then have "?XX x y * log b (?XX x y / (?X x * ?X y)) = - (if x = y then - ?X y * log b (?X y) else 0)" - by (auto simp: log_simps zero_less_mult_iff) } - note remove_XX = this - - show ?thesis - unfolding entropy_def mutual_information_generic_eq[OF MX MX] - unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX - using MX.finite_space by (auto simp: setsum_cases) +lemma (in information_space) entropy_distr: + fixes X :: "'a \ 'b" + assumes "sigma_finite_measure MX" and X: "distributed M MX X f" + shows "entropy b MX X = - (\x. f x * log b (f x) \MX)" +proof - + interpret MX: sigma_finite_measure MX by fact + from X show ?thesis + unfolding entropy_def X[THEN distributed_distr_eq_density] + by (subst MX.KL_density[OF b_gt_1]) (simp_all add: distributed_real_AE distributed_real_measurable) qed -lemma (in information_space) entropy_eq: - assumes "simple_function M X" - shows "\(X) = -(\ x \ X ` space M. distribution X {x} * log b (distribution X {x}))" - using assms by (simp add: entropy_generic_eq) - -lemma (in information_space) entropy_positive: - "simple_function M X \ 0 \ \(X)" - unfolding entropy_def by (simp add: mutual_information_positive) +lemma (in information_space) entropy_uniform: + assumes "sigma_finite_measure MX" + assumes A: "A \ sets MX" "emeasure MX A \ 0" "emeasure MX A \ \" + assumes X: "distributed M MX X (\x. 1 / measure MX A * indicator A x)" + shows "entropy b MX X = log b (measure MX A)" +proof (subst entropy_distr[OF _ X]) + let ?f = "\x. 1 / measure MX A * indicator A x" + have "- (\x. ?f x * log b (?f x) \MX) = + - (\x. (log b (1 / measure MX A) / measure MX A) * indicator A x \MX)" + by (auto intro!: integral_cong simp: indicator_def) + also have "\ = - log b (inverse (measure MX A))" + using A by (subst integral_cmult(2)) + (simp_all add: measure_def real_of_ereal_eq_0 integral_cmult inverse_eq_divide) + also have "\ = log b (measure MX A)" + using b_gt_1 A by (subst log_inverse) (auto simp add: measure_def less_le real_of_ereal_eq_0 + emeasure_nonneg real_of_ereal_pos) + finally show "- (\x. ?f x * log b (?f x) \MX) = log b (measure MX A)" by simp +qed fact+ -lemma (in information_space) entropy_certainty_eq_0: - assumes X: "simple_function M X" and "x \ X ` space M" and "distribution X {x} = 1" - shows "\(X) = 0" -proof - - let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = ereal\distribution X\" - note simple_function_imp_finite_random_variable[OF `simple_function M X`] - from distribution_finite_prob_space[OF this, of "\ measure = ereal\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 - finally have X0: "distribution X (X ` space M - {x}) = 0" by auto - { fix y assume *: "y \ X ` space M" - { assume asm: "y \ x" - with * have "{y} \ X ` space M - {x}" by auto - from X.measure_mono[OF this] X0 asm * - have "distribution X {y} = 0" by (auto intro: antisym) } - then have "distribution X {y} = (if x = y then 1 else 0)" - using assms by auto } - note fi = this - 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 M X`] by (auto simp: y fi) +lemma (in information_space) entropy_simple_distributed: + fixes X :: "'a \ 'b" + assumes X: "simple_distributed M X f" + shows "\(X) = - (\x\X`space M. f x * log b (f x))" +proof (subst entropy_distr[OF _ simple_distributed[OF X]]) + show "sigma_finite_measure (count_space (X ` space M))" + using X by (simp add: sigma_finite_measure_count_space_finite simple_distributed_def) + show "- (\x. f x * log b (f x) \(count_space (X`space M))) = - (\x\X ` space M. f x * log b (f x))" + using X by (auto simp add: lebesgue_integral_count_space_finite) qed lemma (in information_space) entropy_le_card_not_0: - assumes X: "simple_function M X" - shows "\(X) \ log b (card (X ` space M \ {x. distribution X {x} \ 0}))" + assumes X: "simple_distributed M X f" + shows "\(X) \ log b (card (X ` space M \ {x. f x \ 0}))" proof - - let ?p = "\x. distribution X {x}" - have "\(X) = (\x\X`space M. ?p x * log b (1 / ?p x))" - unfolding entropy_eq[OF X] setsum_negf[symmetric] - by (auto intro!: setsum_cong simp: log_simps) - also have "\ \ log b (\x\X`space M. ?p x * (1 / ?p x))" - using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X] - by (intro log_setsum') (auto simp: simple_function_def) - also have "\ = log b (\x\X`space M. if ?p x \ 0 then 1 else 0)" + have "\(X) = (\x\X`space M. f x * log b (1 / f x))" + unfolding entropy_simple_distributed[OF X] setsum_negf[symmetric] + using X by (auto dest: simple_distributed_nonneg intro!: setsum_cong simp: log_simps less_le) + also have "\ \ log b (\x\X`space M. f x * (1 / f x))" + using not_empty b_gt_1 `simple_distributed M X f` + by (intro log_setsum') (auto simp: simple_distributed_nonneg simple_distributed_setsum_space) + also have "\ = log b (\x\X`space M. if f x \ 0 then 1 else 0)" by (intro arg_cong[where f="\X. log b X"] setsum_cong) auto finally show ?thesis - using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def) -qed - -lemma (in prob_space) measure'_translate: - assumes X: "random_variable S X" and A: "A \ sets S" - shows "finite_measure.\' (S\ measure := ereal\distribution X \) A = distribution X A" -proof - - interpret S: prob_space "S\ measure := ereal\distribution X \" - using distribution_prob_space[OF X] . - from A show "S.\' A = distribution X A" - unfolding S.\'_def by (simp add: distribution_def [abs_def] \'_def) -qed - -lemma (in information_space) entropy_uniform_max: - assumes X: "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 - - let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = undefined\\ measure := ereal\distribution X\" - note frv = simple_function_imp_finite_random_variable[OF X] - from distribution_finite_prob_space[OF this, of "\ measure = ereal\distribution X \"] - interpret X: finite_prob_space ?X by simp - note rv = finite_random_variableD[OF frv] - have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff - using `simple_function M X` not_empty by (auto simp: simple_function_def) - { fix x assume "x \ space ?X" - moreover then have "X.\' {x} = 1 / card (space ?X)" - proof (rule X.uniform_prob) - fix x y assume "x \ space ?X" "y \ space ?X" - with assms(2)[of x y] show "X.\' {x} = X.\' {y}" - by (subst (1 2) measure'_translate[OF rv]) auto - qed - ultimately have "distribution X {x} = 1 / card (space ?X)" - by (subst (asm) measure'_translate[OF rv]) auto } - thus ?thesis - using not_empty X.finite_space b_gt_1 card_gt0 - by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps) + using `simple_distributed M X f` by (auto simp: setsum_cases real_eq_of_nat) qed lemma (in information_space) entropy_le_card: - assumes "simple_function M X" + assumes "simple_distributed M X f" 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 + assume "X ` space M \ {x. f x \ 0} = {}" + then have "\x. x\X`space M \ f x = 0" by auto moreover have "0 < card (X`space M)" - using `simple_function M X` not_empty - by (auto simp: card_gt_0_iff simple_function_def) + using `simple_distributed M X f` not_empty by (auto simp: card_gt_0_iff) then have "log b 1 \ log b (real (card (X`space M)))" using b_gt_1 by (intro log_le) auto - ultimately show ?thesis using assms by (simp add: entropy_eq) + ultimately show ?thesis using assms by (simp add: entropy_simple_distributed) next - assume False: "X ` space M \ {x. distribution X {x} \ 0} \ {}" - have "card (X ` space M \ {x. distribution X {x} \ 0}) \ card (X ` space M)" - (is "?A \ ?B") using assms not_empty by (auto intro!: card_mono simp: simple_function_def) + assume False: "X ` space M \ {x. f x \ 0} \ {}" + have "card (X ` space M \ {x. f x \ 0}) \ card (X ` space M)" + (is "?A \ ?B") using assms not_empty + by (auto intro!: card_mono simp: simple_function_def simple_distributed_def) note entropy_le_card_not_0[OF assms] also have "log b (real ?A) \ log b (real ?B)" using b_gt_1 False not_empty `?A \ ?B` assms - by (auto intro!: log_le simp: card_gt_0_iff simp: simple_function_def) + by (auto intro!: log_le simp: card_gt_0_iff simp: simple_distributed_def) finally show ?thesis . qed -lemma (in information_space) entropy_commute: - 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 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 - have inj: "\X. inj_on (\(a,b). (b,a)) X" - by (auto intro!: inj_onI) - show ?thesis - unfolding sf[THEN entropy_eq] unfolding * setsum_reindex[OF inj] - by (simp add: joint_distribution_commute[of Y X] split_beta) -qed - -lemma (in information_space) entropy_eq_cartesian_product: - assumes "simple_function M X" "simple_function M Y" - shows "\(\x. (X x, Y x)) = -(\x\X`space M. \y\Y`space M. - joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))" -proof - - 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 - then have "joint_distribution X Y {x} = 0" - unfolding distribution_def by auto } - then show ?thesis using sf assms - unfolding entropy_eq[OF sf] neg_equal_iff_equal setsum_cartesian_product - by (auto intro!: setsum_mono_zero_cong_left simp: simple_function_def) -qed - subsection {* Conditional Mutual Information *} definition (in prob_space) @@ -917,489 +755,553 @@ 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), measure = ereal\distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = ereal\distribution Y \ - \ space = Z`space M, sets = Pow (Z`space M), measure = ereal\distribution Z \ - X Y Z" + (count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) 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" - and MZ: "finite_random_variable MZ Z" - shows "conditional_mutual_information b MX MY MZ X Y Z = (\(x, y, z) \ space MX \ space MY \ space MZ. - distribution (\x. (X x, Y x, Z x)) {(x, y, z)} * - log b (distribution (\x. (X x, Y x, Z x)) {(x, y, z)} / - (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" - (is "_ = (\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))") + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" + assumes Px: "distributed M S X Px" + assumes Pz: "distributed M P Z Pz" + assumes Pyz: "distributed M (T \\<^isub>M P) (\x. (Y x, Z x)) Pyz" + assumes Pxz: "distributed M (S \\<^isub>M P) (\x. (X x, Z x)) Pxz" + assumes Pxyz: "distributed M (S \\<^isub>M T \\<^isub>M P) (\x. (X x, Y x, Z x)) Pxyz" + assumes I1: "integrable (S \\<^isub>M T \\<^isub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))" + assumes I2: "integrable (S \\<^isub>M T \\<^isub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))" + shows "conditional_mutual_information b S T P X Y Z + = (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \(S \\<^isub>M T \\<^isub>M P))" proof - - let ?X = "\x. distribution X {x}" - note finite_var = MX MY MZ - note YZ = finite_random_variable_pairI[OF finite_var(2,3)] - note XYZ = finite_random_variable_pairI[OF MX YZ] - note XZ = finite_random_variable_pairI[OF finite_var(1,3)] - note ZX = finite_random_variable_pairI[OF finite_var(3,1)] - note YZX = finite_random_variable_pairI[OF finite_var(2) ZX] - note order1 = - finite_distribution_order(5,6)[OF finite_var(1) YZ] - finite_distribution_order(5,6)[OF finite_var(1,3)] + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret P: sigma_finite_measure P by fact + interpret TP: pair_sigma_finite T P .. + interpret SP: pair_sigma_finite S P .. + interpret SPT: pair_sigma_finite "S \\<^isub>M P" T .. + interpret STP: pair_sigma_finite S "T \\<^isub>M P" .. + have TP: "sigma_finite_measure (T \\<^isub>M P)" .. + have SP: "sigma_finite_measure (S \\<^isub>M P)" .. + have YZ: "random_variable (T \\<^isub>M P) (\x. (Y x, Z x))" + using Pyz by (simp add: distributed_measurable) + + have Pxyz_f: "\M f. f \ measurable M (S \\<^isub>M T \\<^isub>M P) \ (\x. Pxyz (f x)) \ borel_measurable M" + using measurable_comp[OF _ Pxyz[THEN distributed_real_measurable]] by (auto simp: comp_def) + + { fix f g h M + assume f: "f \ measurable M S" and g: "g \ measurable M P" and h: "h \ measurable M (S \\<^isub>M P)" + from measurable_comp[OF h Pxz[THEN distributed_real_measurable]] + measurable_comp[OF f Px[THEN distributed_real_measurable]] + measurable_comp[OF g Pz[THEN distributed_real_measurable]] + have "(\x. log b (Pxz (h x) / (Px (f x) * Pz (g x)))) \ borel_measurable M" + by (simp add: comp_def b_gt_1) } + note borel_log = this + + have measurable_cut: "(\(x, y, z). (x, z)) \ measurable (S \\<^isub>M T \\<^isub>M P) (S \\<^isub>M P)" + by (auto simp add: split_beta' comp_def intro!: measurable_Pair measurable_snd') + + from Pxz Pxyz have distr_eq: "distr M (S \\<^isub>M P) (\x. (X x, Z x)) = + distr (distr M (S \\<^isub>M T \\<^isub>M P) (\x. (X x, Y x, Z x))) (S \\<^isub>M P) (\(x, y, z). (x, z))" + by (subst distr_distr[OF measurable_cut]) (auto dest: distributed_measurable simp: comp_def) - note random_var = finite_var[THEN finite_random_variableD] - note finite = finite_var(1) YZ finite_var(3) XZ YZX - - have order2: "\x y z. \x \ space MX; y \ space MY; z \ space MZ; joint_distribution X Z {(x, z)} = 0\ - \ joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = 0" - unfolding joint_distribution_commute_singleton[of X] - unfolding joint_distribution_assoc_singleton[symmetric] - using finite_distribution_order(6)[OF finite_var(2) ZX] - by auto - - have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z)))) = - (\(x, y, z)\?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))" - (is "(\(x, y, z)\?S. ?L x y z) = (\(x, y, z)\?S. ?R x y z)") - proof (safe intro!: setsum_cong) - fix x y z assume space: "x \ space MX" "y \ space MY" "z \ space MZ" - show "?L x y z = ?R x y z" + have "mutual_information b S P X Z = + (\x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \(S \\<^isub>M P))" + by (rule mutual_information_distr[OF S P Px Pz Pxz]) + also have "\ = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^isub>M T \\<^isub>M P))" + using b_gt_1 Pxz Px Pz + by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\(x, y, z). (x, z)"]) + (auto simp: split_beta' intro!: measurable_Pair measurable_snd' measurable_snd'' measurable_fst'' borel_measurable_times + dest!: distributed_real_measurable) + finally have mi_eq: + "mutual_information b S P X Z = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^isub>M T \\<^isub>M P))" . + + have "AE x in S \\<^isub>M T \\<^isub>M P. Px (fst x) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. Pz (snd (snd x)) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of "\x. snd (snd x)", OF _ Pxyz Pz]) (auto intro: measurable_snd') + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of "\x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto intro: measurable_Pair measurable_snd') + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. Pyz (snd x) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto intro: measurable_Pair) + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. 0 \ Px (fst x)" + using Px by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable) + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. 0 \ Pyz (snd x)" + using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable) + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. 0 \ Pz (snd (snd x))" + using Pz Pz[THEN distributed_real_measurable] by (auto intro!: measurable_snd'' TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE) + moreover have "AE x in S \\<^isub>M T \\<^isub>M P. 0 \ Pxz (fst x, snd (snd x))" + using Pxz[THEN distributed_real_AE, THEN SP.AE_pair] + using measurable_comp[OF measurable_Pair[OF measurable_fst measurable_comp[OF measurable_snd measurable_snd]] Pxz[THEN distributed_real_measurable], of T] + using measurable_comp[OF measurable_snd measurable_Pair2[OF Pxz[THEN distributed_real_measurable]], of _ T] + by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def) + moreover note Pxyz[THEN distributed_real_AE] + ultimately have "AE x in S \\<^isub>M T \\<^isub>M P. + Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) - + Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) = + Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) " + proof eventually_elim + case (goal1 x) + show ?case proof cases - assume "?XYZ x y z \ 0" - with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z" - using order1 order2 by (auto simp: less_le) - with b_gt_1 show ?thesis - by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff) + assume "Pxyz x \ 0" + with goal1 have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" "0 < Pyz (snd x)" "0 < Pxyz x" + by auto + then show ?thesis + using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps) qed simp qed - also have "\ = (\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - - (\(x, y, z)\?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z)))" - by (auto simp add: setsum_subtractf[symmetric] field_simps intro!: setsum_cong) - also have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z))) = - (\(x, z)\space MX \ space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z)))" - unfolding setsum_cartesian_product[symmetric] setsum_commute[of _ _ "space MY"] - setsum_left_distrib[symmetric] - unfolding joint_distribution_commute_singleton[of X] - unfolding joint_distribution_assoc_singleton[symmetric] - using setsum_joint_distribution_singleton[OF finite_var(2) ZX] - by (intro setsum_cong refl) (simp add: space_pair_measure) - also have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - - (\(x, z)\space MX \ space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) = - conditional_mutual_information b MX MY MZ X Y Z" + with I1 I2 show ?thesis 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_measure setsum_cartesian_product') - finally show ?thesis by simp + apply (subst mi_eq) + apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz]) + apply (subst integral_diff(2)[symmetric]) + apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff) + done qed lemma (in information_space) conditional_mutual_information_eq: - 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. - distribution (\x. (X x, Y x, Z x)) {(x, y, z)} * - log b (distribution (\x. (X x, Y x, Z x)) {(x, y, z)} / - (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))" - by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) - simp + assumes Pz: "simple_distributed M Z Pz" + assumes Pyz: "simple_distributed M (\x. (Y x, Z x)) Pyz" + assumes Pxz: "simple_distributed M (\x. (X x, Z x)) Pxz" + assumes Pxyz: "simple_distributed M (\x. (X x, Y x, Z x)) Pxyz" + shows "\(X ; Y | Z) = + (\(x, y, z)\(\x. (X x, Y x, Z x))`space M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" +proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _ + simple_distributed[OF Pz] simple_distributed_joint[OF Pyz] simple_distributed_joint[OF Pxz] + simple_distributed_joint2[OF Pxyz]]) + note simple_distributed_joint2_finite[OF Pxyz, simp] + show "sigma_finite_measure (count_space (X ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + show "sigma_finite_measure (count_space (Y ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + show "sigma_finite_measure (count_space (Z ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + have "count_space (X ` space M) \\<^isub>M count_space (Y ` space M) \\<^isub>M count_space (Z ` space M) = + count_space (X`space M \ Y`space M \ Z`space M)" + (is "?P = ?C") + by (simp add: pair_measure_count_space) -lemma (in information_space) conditional_mutual_information_eq_mutual_information: - 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 M (\x. ())" by auto - show ?thesis - unfolding conditional_mutual_information_eq[OF X Y C] - unfolding mutual_information_eq[OF X Y] - by (simp add: setsum_cartesian_product' distribution_remove_const) + let ?Px = "\x. measure M (X -` {x} \ space M)" + have "(\x. (X x, Z x)) \ measurable M (count_space (X ` space M) \\<^isub>M count_space (Z ` space M))" + using simple_distributed_joint[OF Pxz] by (rule distributed_measurable) + from measurable_comp[OF this measurable_fst] + have "random_variable (count_space (X ` space M)) X" + by (simp add: comp_def) + then have "simple_function M X" + unfolding simple_function_def by auto + then have "simple_distributed M X ?Px" + by (rule simple_distributedI) auto + then show "distributed M (count_space (X ` space M)) X ?Px" + by (rule simple_distributed) + + let ?f = "(\x. if x \ (\x. (X x, Y x, Z x)) ` space M then Pxyz x else 0)" + let ?g = "(\x. if x \ (\x. (Y x, Z x)) ` space M then Pyz x else 0)" + let ?h = "(\x. if x \ (\x. (X x, Z x)) ` space M then Pxz x else 0)" + show + "integrable ?P (\(x, y, z). ?f (x, y, z) * log b (?f (x, y, z) / (?Px x * ?g (y, z))))" + "integrable ?P (\(x, y, z). ?f (x, y, z) * log b (?h (x, z) / (?Px x * Pz z)))" + by (auto intro!: integrable_count_space simp: pair_measure_count_space) + let ?i = "\x y z. ?f (x, y, z) * log b (?f (x, y, z) / (?h (x, z) * (?g (y, z) / Pz z)))" + let ?j = "\x y z. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z)))" + have "(\(x, y, z). ?i x y z) = (\x. if x \ (\x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)" + by (auto intro!: ext) + then show "(\ (x, y, z). ?i x y z \?P) = (\(x, y, z)\(\x. (X x, Y x, Z x)) ` space M. ?j x y z)" + by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite setsum_cases split_beta') qed -lemma (in information_space) conditional_mutual_information_generic_positive: - assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "finite_random_variable MZ Z" - shows "0 \ conditional_mutual_information b MX MY MZ X Y Z" -proof (cases "space MX \ space MY \ space MZ = {}") - case True show ?thesis - unfolding conditional_mutual_information_generic_eq[OF assms] True - by simp -next - case False - let ?dXYZ = "distribution (\x. (X x, Y x, Z x))" - let ?dXZ = "joint_distribution X Z" - let ?dYZ = "joint_distribution Y Z" - let ?dX = "distribution X" - let ?dZ = "distribution Z" - let ?M = "space MX \ space MY \ space MZ" +lemma (in information_space) conditional_mutual_information_nonneg: + assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z" + shows "0 \ \(X ; Y | Z)" +proof - + def Pz \ "\x. if x \ Z`space M then measure M (Z -` {x} \ space M) else 0" + def Pxz \ "\x. if x \ (\x. (X x, Z x))`space M then measure M ((\x. (X x, Z x)) -` {x} \ space M) else 0" + def Pyz \ "\x. if x \ (\x. (Y x, Z x))`space M then measure M ((\x. (Y x, Z x)) -` {x} \ space M) else 0" + def Pxyz \ "\x. if x \ (\x. (X x, Y x, Z x))`space M then measure M ((\x. (X x, Y x, Z x)) -` {x} \ space M) else 0" + let ?M = "X`space M \ Y`space M \ Z`space M" - note YZ = finite_random_variable_pairI[OF Y Z] - note XZ = finite_random_variable_pairI[OF X Z] - note ZX = finite_random_variable_pairI[OF Z X] - note YZ = finite_random_variable_pairI[OF Y Z] - note XYZ = finite_random_variable_pairI[OF X YZ] - note finite = Z YZ XZ XYZ - have order: "\x y z. \x \ space MX; y \ space MY; z \ space MZ; joint_distribution X Z {(x, z)} = 0\ - \ joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = 0" - unfolding joint_distribution_commute_singleton[of X] - unfolding joint_distribution_assoc_singleton[symmetric] - using finite_distribution_order(6)[OF Y ZX] - by auto + note XZ = simple_function_Pair[OF X Z] + note YZ = simple_function_Pair[OF Y Z] + note XYZ = simple_function_Pair[OF X simple_function_Pair[OF Y Z]] + have Pz: "simple_distributed M Z Pz" + using Z by (rule simple_distributedI) (auto simp: Pz_def) + have Pxz: "simple_distributed M (\x. (X x, Z x)) Pxz" + using XZ by (rule simple_distributedI) (auto simp: Pxz_def) + have Pyz: "simple_distributed M (\x. (Y x, Z x)) Pyz" + using YZ by (rule simple_distributedI) (auto simp: Pyz_def) + have Pxyz: "simple_distributed M (\x. (X x, Y x, Z x)) Pxyz" + using XYZ by (rule simple_distributedI) (auto simp: Pxyz_def) - note order = order - finite_distribution_order(5,6)[OF X YZ] - finite_distribution_order(5,6)[OF Y Z] + { fix z assume z: "z \ Z ` space M" then have "(\x\X ` space M. Pxz (x, z)) = Pz z" + using distributed_marginal_eq_joint_simple[OF X Pz Pxz z] + by (auto intro!: setsum_cong simp: Pxz_def) } + note marginal1 = this - have "- conditional_mutual_information b MX MY MZ X Y Z = - (\(x, y, z) \ ?M. ?dXYZ {(x, y, z)} * - log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))" - unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal by auto - also have "\ \ log b (\(x, y, z) \ ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})" + { fix z assume z: "z \ Z ` space M" then have "(\y\Y ` space M. Pyz (y, z)) = Pz z" + using distributed_marginal_eq_joint_simple[OF Y Pz Pyz z] + by (auto intro!: setsum_cong simp: Pyz_def) } + note marginal2 = this + + have "- \(X ; Y | Z) = - (\(x, y, z) \ ?M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" + unfolding conditional_mutual_information_eq[OF Pz Pyz Pxz Pxyz] + using X Y Z by (auto intro!: setsum_mono_zero_left simp: Pxyz_def simple_functionD) + also have "\ \ log b (\(x, y, z) \ ?M. Pxz (x, z) * (Pyz (y,z) / Pz z))" unfolding split_beta' proof (rule log_setsum_divide) - show "?M \ {}" using False by simp + show "?M \ {}" using not_empty by simp show "1 < b" using b_gt_1 . - show "finite ?M" using assms - unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by auto - - show "(\x\?M. ?dXYZ {(fst x, fst (snd x), snd (snd x))}) = 1" - unfolding setsum_cartesian_product' - unfolding setsum_commute[of _ "space MY"] - unfolding setsum_commute[of _ "space MZ"] - by (simp_all add: space_pair_measure - setsum_joint_distribution_singleton[OF X YZ] - setsum_joint_distribution_singleton[OF Y Z] - setsum_distribution[OF Z]) + show "finite ?M" using X Y Z by (auto simp: simple_functionD) - fix x assume "x \ ?M" - let ?x = "(fst x, fst (snd x), snd (snd x))" - - show "0 \ ?dXYZ {?x}" - "0 \ ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" - by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg) - - assume *: "0 < ?dXYZ {?x}" - with `x \ ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" - by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le) + then show "(\x\?M. Pxyz (fst x, fst (snd x), snd (snd x))) = 1" + apply (subst Pxyz[THEN simple_distributed_setsum_space, symmetric]) + apply simp + apply (intro setsum_mono_zero_right) + apply (auto simp: Pxyz_def) + done + let ?N = "(\x. (X x, Y x, Z x)) ` space M" + fix x assume x: "x \ ?M" + let ?Q = "Pxyz (fst x, fst (snd x), snd (snd x))" + let ?P = "Pxz (fst x, snd (snd x)) * (Pyz (fst (snd x), snd (snd x)) / Pz (snd (snd x)))" + from x show "0 \ ?Q" "0 \ ?P" + using Pxyz[THEN simple_distributed, THEN distributed_real_AE] + using Pxz[THEN simple_distributed, THEN distributed_real_AE] + using Pyz[THEN simple_distributed, THEN distributed_real_AE] + using Pz[THEN simple_distributed, THEN distributed_real_AE] + by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg simp: AE_count_space Pxyz_def Pxz_def Pyz_def Pz_def) + moreover assume "0 < ?Q" + moreover have "AE x in count_space ?N. Pz (snd (snd x)) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of "\x. snd (snd x)", OF _ Pxyz[THEN simple_distributed] Pz[THEN simple_distributed]]) (auto intro: measurable_snd') + then have "\x. Pz (snd (snd x)) = 0 \ Pxyz x = 0" + by (auto simp: Pz_def Pxyz_def AE_count_space) + moreover have "AE x in count_space ?N. Pxz (fst x, snd (snd x)) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of "\x. (fst x, snd (snd x))", OF _ Pxyz[THEN simple_distributed] Pxz[THEN simple_distributed]]) (auto intro: measurable_Pair measurable_snd') + then have "\x. Pxz (fst x, snd (snd x)) = 0 \ Pxyz x = 0" + by (auto simp: Pz_def Pxyz_def AE_count_space) + moreover have "AE x in count_space ?N. Pyz (snd x) = 0 \ Pxyz x = 0" + by (intro subdensity_real[of snd, OF _ Pxyz[THEN simple_distributed] Pyz[THEN simple_distributed]]) (auto intro: measurable_Pair) + then have "\x. Pyz (snd x) = 0 \ Pxyz x = 0" + by (auto simp: Pz_def Pxyz_def AE_count_space) + ultimately show "0 < ?P" using x by (auto intro!: divide_pos_pos mult_pos_pos simp: less_le) qed - also have "(\(x, y, z) \ ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\z\space MZ. ?dZ {z})" + also have "(\(x, y, z) \ ?M. Pxz (x, z) * (Pyz (y,z) / Pz z)) = (\z\Z`space M. Pz z)" apply (simp add: setsum_cartesian_product') apply (subst setsum_commute) apply (subst (2) setsum_commute) - by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] - setsum_joint_distribution_singleton[OF X Z] - setsum_joint_distribution_singleton[OF Y Z] + apply (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] marginal1 marginal2 intro!: setsum_cong) - also have "log b (\z\space MZ. ?dZ {z}) = 0" - unfolding setsum_distribution[OF Z] by simp + done + also have "log b (\z\Z`space M. Pz z) = 0" + using Pz[THEN simple_distributed_setsum_space] by simp finally show ?thesis by simp qed -lemma (in information_space) conditional_mutual_information_positive: - assumes "simple_function M X" and "simple_function M Y" and "simple_function M Z" - shows "0 \ \(X;Y|Z)" - by (rule conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]]) - subsection {* Conditional Entropy *} definition (in prob_space) - "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y" + "conditional_entropy b S T X Y = entropy b (S \\<^isub>M T) (\x. (X x, Y x)) - entropy b T Y" abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where - "\(X | Y) \ conditional_entropy b - \ space = X`space M, sets = Pow (X`space M), measure = ereal\distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = ereal\distribution Y \ X Y" - -lemma (in information_space) conditional_entropy_positive: - "simple_function M X \ simple_function M Y \ 0 \ \(X | Y)" - unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive) + "\(X | Y) \ conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y" 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 = - - (\(x, z)\space MX \ space MZ. - joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))" + fixes Px :: "'b \ real" and Py :: "'c \ real" + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes Px: "distributed M S X Px" + assumes Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + assumes I1: "integrable (S \\<^isub>M T) (\x. Pxy x * log b (Pxy x))" + assumes I2: "integrable (S \\<^isub>M T) (\x. Pxy x * log b (Py (snd x)))" + shows "conditional_entropy b S T X Y = - (\(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \(S \\<^isub>M T))" proof - - interpret MX: finite_sigma_algebra MX using MX by simp - interpret MZ: finite_sigma_algebra MZ using MZ by simp - let ?XXZ = "\x y z. joint_distribution X (\x. (X x, Z x)) {(x, y, z)}" - let ?XZ = "\x z. joint_distribution X Z {(x, z)}" - let ?Z = "\z. distribution Z {z}" - let ?f = "\x y z. log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y z))" - { 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' :: 'c and z assume "x' \ x" - then have "?XXZ x x' z = 0" - by (auto simp: distribution_def empty_measure'[symmetric] - simp del: empty_measure' intro!: arg_cong[where f=\']) } - note this[simp] - { fix x x' z assume *: "x \ space MX" "z \ space MZ" - then have "(\x'\space MX. ?XXZ x x' z * ?f x x' z) - = (\x'\space MX. if x = x' then ?XZ x z * ?f x x z else 0)" - by (auto intro!: setsum_cong) - also have "\ = ?XZ x z * ?f x x z" - using `x \ space MX` by (simp add: setsum_cases[OF MX.finite_space]) - also have "\ = ?XZ x z * log b (?Z z / ?XZ x z)" by auto - also have "\ = - ?XZ x z * log b (?XZ x z / ?Z z)" - using finite_distribution_order(6)[OF MX MZ] - by (auto simp: log_simps field_simps zero_less_mult_iff) - finally have "(\x'\space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . } - note * = this - show ?thesis + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T .. + have ST: "sigma_finite_measure (S \\<^isub>M T)" .. + + interpret Pxy: prob_space "density (S \\<^isub>M T) Pxy" + unfolding Pxy[THEN distributed_distr_eq_density, symmetric] + using Pxy[THEN distributed_measurable] by (rule prob_space_distr) + + from Py Pxy have distr_eq: "distr M T Y = + distr (distr M (S \\<^isub>M T) (\x. (X x, Y x))) T snd" + by (subst distr_distr[OF measurable_snd]) (auto dest: distributed_measurable simp: comp_def) + + have "entropy b T Y = - (\y. Py y * log b (Py y) \T)" + by (rule entropy_distr[OF T Py]) + also have "\ = - (\(x,y). Pxy (x,y) * log b (Py y) \(S \\<^isub>M T))" + using b_gt_1 Py[THEN distributed_real_measurable] + by (subst distributed_transform_integral[OF Pxy Py, where T=snd]) (auto intro!: integral_cong) + finally have e_eq: "entropy b T Y = - (\(x,y). Pxy (x,y) * log b (Py y) \(S \\<^isub>M T))" . + + have "AE x in S \\<^isub>M T. Px (fst x) = 0 \ Pxy x = 0" + by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair) + moreover have "AE x in S \\<^isub>M T. Py (snd x) = 0 \ Pxy x = 0" + by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair) + moreover have "AE x in S \\<^isub>M T. 0 \ Px (fst x)" + using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable) + moreover have "AE x in S \\<^isub>M T. 0 \ Py (snd x)" + using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable) + moreover note Pxy[THEN distributed_real_AE] + ultimately have pos: "AE x in S \\<^isub>M T. 0 \ Pxy x \ 0 \ Px (fst x) \ 0 \ Py (snd x) \ + (Pxy x = 0 \ (Pxy x \ 0 \ 0 < Pxy x \ 0 < Px (fst x) \ 0 < Py (snd x)))" + by eventually_elim auto + + from pos have "AE x in S \\<^isub>M T. + Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))" + by eventually_elim (auto simp: log_simps mult_pos_pos field_simps b_gt_1) + with I1 I2 show ?thesis unfolding conditional_entropy_def - unfolding conditional_mutual_information_generic_eq[OF MX MX MZ] - by (auto simp: setsum_cartesian_product' setsum_negf[symmetric] - setsum_commute[of _ "space MZ"] * - intro!: setsum_cong) + apply (subst e_eq) + apply (subst entropy_distr[OF ST Pxy]) + unfolding minus_diff_minus + apply (subst integral_diff(2)[symmetric]) + apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff) + done qed lemma (in information_space) conditional_entropy_eq: - assumes "simple_function M X" "simple_function M Z" - shows "\(X | Z) = - - (\(x, z)\X ` space M \ Z ` space M. - joint_distribution X Z {(x, z)} * - log b (joint_distribution X Z {(x, z)} / distribution Z {z}))" - by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) - simp + assumes Y: "simple_distributed M Y Py" and X: "simple_function M X" + assumes XY: "simple_distributed M (\x. (X x, Y x)) Pxy" + shows "\(X | Y) = - (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" +proof (subst conditional_entropy_generic_eq[OF _ _ + simple_distributed[OF simple_distributedI[OF X refl]] simple_distributed[OF Y] simple_distributed_joint[OF XY]]) + have [simp]: "finite (X`space M)" using X by (simp add: simple_function_def) + note Y[THEN simple_distributed_finite, simp] + show "sigma_finite_measure (count_space (X ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + show "sigma_finite_measure (count_space (Y ` space M))" + by (simp add: sigma_finite_measure_count_space_finite) + let ?f = "(\x. if x \ (\x. (X x, Y x)) ` space M then Pxy x else 0)" + have "count_space (X ` space M) \\<^isub>M count_space (Y ` space M) = count_space (X`space M \ Y`space M)" + (is "?P = ?C") + using X Y by (simp add: simple_distributed_finite pair_measure_count_space) + with X Y show + "integrable ?P (\x. ?f x * log b (?f x))" + "integrable ?P (\x. ?f x * log b (Py (snd x)))" + by (auto intro!: integrable_count_space simp: simple_distributed_finite) + have eq: "(\(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) = + (\x. if x \ (\x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / Py (snd x)) else 0)" + by auto + from X Y show "- (\ (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \?P) = + - (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" + by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum_cases split_beta') +qed -lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis: +lemma (in information_space) conditional_mutual_information_eq_conditional_entropy: assumes X: "simple_function M X" and Y: "simple_function M Y" - shows "\(X | Y) = - -(\y\Y`space M. distribution Y {y} * - (\x\X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} * - log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))" - unfolding conditional_entropy_eq[OF assms] - using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]] - by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib - intro!: setsum_cong) + shows "\(X ; X | Y) = \(X | Y)" +proof - + def Py \ "\x. if x \ Y`space M then measure M (Y -` {x} \ space M) else 0" + def Pxy \ "\x. if x \ (\x. (X x, Y x))`space M then measure M ((\x. (X x, Y x)) -` {x} \ space M) else 0" + def Pxxy \ "\x. if x \ (\x. (X x, X x, Y x))`space M then measure M ((\x. (X x, X x, Y x)) -` {x} \ space M) else 0" + let ?M = "X`space M \ X`space M \ Y`space M" -lemma (in information_space) conditional_entropy_eq_cartesian_product: - assumes "simple_function M X" "simple_function M Y" - shows "\(X | Y) = -(\x\X`space M. \y\Y`space M. - joint_distribution X Y {(x,y)} * - log b (joint_distribution X Y {(x,y)} / distribution Y {y}))" - unfolding conditional_entropy_eq[OF assms] - by (auto intro!: setsum_cong simp: setsum_cartesian_product') + note XY = simple_function_Pair[OF X Y] + note XXY = simple_function_Pair[OF X XY] + have Py: "simple_distributed M Y Py" + using Y by (rule simple_distributedI) (auto simp: Py_def) + have Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" + using XY by (rule simple_distributedI) (auto simp: Pxy_def) + have Pxxy: "simple_distributed M (\x. (X x, X x, Y x)) Pxxy" + using XXY by (rule simple_distributedI) (auto simp: Pxxy_def) + have eq: "(\x. (X x, X x, Y x)) ` space M = (\(x, y). (x, x, y)) ` (\x. (X x, Y x)) ` space M" + by auto + have inj: "\A. inj_on (\(x, y). (x, x, y)) A" + by (auto simp: inj_on_def) + have Pxxy_eq: "\x y. Pxxy (x, x, y) = Pxy (x, y)" + by (auto simp: Pxxy_def Pxy_def intro!: arg_cong[where f=prob]) + have "AE x in count_space ((\x. (X x, Y x))`space M). Py (snd x) = 0 \ Pxy x = 0" + by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) (auto intro: measurable_Pair) + then show ?thesis + apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy]) + apply (subst conditional_entropy_eq[OF Py X Pxy]) + apply (auto intro!: setsum_cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum_reindex[OF inj] + log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space) + using Py[THEN simple_distributed, THEN distributed_real_AE] Pxy[THEN simple_distributed, THEN distributed_real_AE] + apply (auto simp add: not_le[symmetric] AE_count_space) + done +qed + +lemma (in information_space) conditional_entropy_nonneg: + assumes X: "simple_function M X" and Y: "simple_function M Y" shows "0 \ \(X | Y)" + using conditional_mutual_information_eq_conditional_entropy[OF X Y] conditional_mutual_information_nonneg[OF X X Y] + by simp subsection {* Equalities *} -lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: - assumes X: "simple_function M X" and Z: "simple_function M Z" - shows "\(X ; Z) = \(X) - \(X | Z)" +lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr: + fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + assumes Ix: "integrable(S \\<^isub>M T) (\x. Pxy x * log b (Px (fst x)))" + assumes Iy: "integrable(S \\<^isub>M T) (\x. Pxy x * log b (Py (snd x)))" + assumes Ixy: "integrable(S \\<^isub>M T) (\x. Pxy x * log b (Pxy x))" + shows "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \\<^isub>M T) (\x. (X x, Y x))" proof - - let ?XZ = "\x z. joint_distribution X Z {(x, z)}" - let ?Z = "\z. distribution Z {z}" - let ?X = "\x. distribution X {x}" - note fX = X[THEN simple_function_imp_finite_random_variable] - note fZ = Z[THEN simple_function_imp_finite_random_variable] - note finite_distribution_order[OF fX fZ, simp] - { fix x z assume "x \ X`space M" "z \ Z`space M" - have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) = - ?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)" - by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } - note * = this + have X: "entropy b S X = - (\x. Pxy x * log b (Px (fst x)) \(S \\<^isub>M T))" + using b_gt_1 Px[THEN distributed_real_measurable] + apply (subst entropy_distr[OF S Px]) + apply (subst distributed_transform_integral[OF Pxy Px, where T=fst]) + apply (auto intro!: integral_cong) + done + + have Y: "entropy b T Y = - (\x. Pxy x * log b (Py (snd x)) \(S \\<^isub>M T))" + using b_gt_1 Py[THEN distributed_real_measurable] + apply (subst entropy_distr[OF T Py]) + apply (subst distributed_transform_integral[OF Pxy Py, where T=snd]) + apply (auto intro!: integral_cong) + done + + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T .. + have ST: "sigma_finite_measure (S \\<^isub>M T)" .. + + have XY: "entropy b (S \\<^isub>M T) (\x. (X x, Y x)) = - (\x. Pxy x * log b (Pxy x) \(S \\<^isub>M T))" + by (subst entropy_distr[OF ST Pxy]) (auto intro!: integral_cong) + + have "AE x in S \\<^isub>M T. Px (fst x) = 0 \ Pxy x = 0" + by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair) + moreover have "AE x in S \\<^isub>M T. Py (snd x) = 0 \ Pxy x = 0" + by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair) + moreover have "AE x in S \\<^isub>M T. 0 \ Px (fst x)" + using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable) + moreover have "AE x in S \\<^isub>M T. 0 \ Py (snd x)" + using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable) + moreover note Pxy[THEN distributed_real_AE] + ultimately have "AE x in S \\<^isub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) = + Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" + (is "AE x in _. ?f x = ?g x") + proof eventually_elim + case (goal1 x) + show ?case + proof cases + assume "Pxy x \ 0" + with goal1 have "0 < Px (fst x)" "0 < Py (snd x)" "0 < Pxy x" + by auto + then show ?thesis + using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps) + qed simp + qed + + have "entropy b S X + entropy b T Y - entropy b (S \\<^isub>M T) (\x. (X x, Y x)) = integral\<^isup>L (S \\<^isub>M T) ?f" + unfolding X Y XY + apply (subst integral_diff) + apply (intro integral_diff Ixy Ix Iy)+ + apply (subst integral_diff) + apply (intro integral_diff Ixy Ix Iy)+ + apply (simp add: field_simps) + done + also have "\ = integral\<^isup>L (S \\<^isub>M T) ?g" + using `AE x in _. ?f x = ?g x` by (rule integral_cong_AE) + also have "\ = mutual_information b S T X Y" + by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric]) + finally show ?thesis .. +qed + +lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: + assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" + shows "\(X ; Y) = \(X) - \(X | Y)" +proof - + have X: "simple_distributed M X (\x. measure M (X -` {x} \ space M))" + using sf_X by (rule simple_distributedI) auto + have Y: "simple_distributed M Y (\x. measure M (Y -` {x} \ space M))" + using sf_Y by (rule simple_distributedI) auto + have sf_XY: "simple_function M (\x. (X x, Y x))" + using sf_X sf_Y by (rule simple_function_Pair) + then have XY: "simple_distributed M (\x. (X x, Y x)) (\x. measure M ((\x. (X x, Y x)) -` {x} \ space M))" + by (rule simple_distributedI) auto + from simple_distributed_joint_finite[OF this, simp] + have eq: "count_space (X ` space M) \\<^isub>M count_space (Y ` space M) = count_space (X ` space M \ Y ` space M)" + by (simp add: pair_measure_count_space) + + have "\(X ; Y) = \(X) + \(Y) - entropy b (count_space (X`space M) \\<^isub>M count_space (Y`space M)) (\x. (X x, Y x))" + using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY] + by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space) + then show ?thesis + unfolding conditional_entropy_def by simp +qed + +lemma (in information_space) mutual_information_nonneg_simple: + assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" + shows "0 \ \(X ; Y)" +proof - + have X: "simple_distributed M X (\x. measure M (X -` {x} \ space M))" + using sf_X by (rule simple_distributedI) auto + have Y: "simple_distributed M Y (\x. measure M (Y -` {x} \ space M))" + using sf_Y by (rule simple_distributedI) auto + + have sf_XY: "simple_function M (\x. (X x, Y x))" + using sf_X sf_Y by (rule simple_function_Pair) + then have XY: "simple_distributed M (\x. (X x, Y x)) (\x. measure M ((\x. (X x, Y x)) -` {x} \ space M))" + by (rule simple_distributedI) auto + + from simple_distributed_joint_finite[OF this, simp] + have eq: "count_space (X ` space M) \\<^isub>M count_space (Y ` space M) = count_space (X ` space M \ Y ` space M)" + by (simp add: pair_measure_count_space) + show ?thesis - unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z] - using setsum_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]] - by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric] - setsum_distribution) + by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]]) + (simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite) qed lemma (in information_space) conditional_entropy_less_eq_entropy: 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] . - with mutual_information_positive[OF X Z] entropy_positive[OF X] - show ?thesis by auto + have "0 \ \(X ; Z)" using X Z by (rule mutual_information_nonneg_simple) + also have "\(X ; Z) = \(X) - \(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] . + finally show ?thesis by auto qed lemma (in information_space) entropy_chain_rule: 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. joint_distribution X Y {(x, y)}" - let ?Y = "\y. distribution Y {y}" - let ?X = "\x. distribution X {x}" - note fX = X[THEN simple_function_imp_finite_random_variable] - note fY = Y[THEN simple_function_imp_finite_random_variable] - note finite_distribution_order[OF fX fY, simp] - { fix x y assume "x \ X`space M" "y \ Y`space M" - have "?XY x y * log b (?XY x y / ?X x) = - ?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)" - by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } - note * = this - show ?thesis - using setsum_joint_distribution_singleton[OF fY fX] - unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y] - unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"] - by (simp add: * setsum_subtractf setsum_left_distrib[symmetric]) -qed - -section {* Partitioning *} - -definition "subvimage A f g \ (\x \ A. f -` {f x} \ A \ g -` {g x} \ A)" - -lemma subvimageI: - assumes "\x y. \ x \ A ; y \ A ; f x = f y \ \ g x = g y" - shows "subvimage A f g" - using assms unfolding subvimage_def by blast - -lemma subvimageE[consumes 1]: - assumes "subvimage A f g" - obtains "\x y. \ x \ A ; y \ A ; f x = f y \ \ g x = g y" - using assms unfolding subvimage_def by blast - -lemma subvimageD: - "\ subvimage A f g ; x \ A ; y \ A ; f x = f y \ \ g x = g y" - using assms unfolding subvimage_def by blast - -lemma subvimage_subset: - "\ subvimage B f g ; A \ B \ \ subvimage A f g" - unfolding subvimage_def by auto - -lemma subvimage_idem[intro]: "subvimage A g g" - by (safe intro!: subvimageI) - -lemma subvimage_comp_finer[intro]: - assumes svi: "subvimage A g h" - shows "subvimage A g (f \ h)" -proof (rule subvimageI, simp) - fix x y assume "x \ A" "y \ A" "g x = g y" - from svi[THEN subvimageD, OF this] - show "f (h x) = f (h y)" by simp -qed - -lemma subvimage_comp_gran: - assumes svi: "subvimage A g h" - assumes inj: "inj_on f (g ` A)" - shows "subvimage A (f \ g) h" - by (rule subvimageI) (auto intro!: subvimageD[OF svi] simp: inj_on_iff[OF inj]) - -lemma subvimage_comp: - assumes svi: "subvimage (f ` A) g h" - shows "subvimage A (g \ f) (h \ f)" - by (rule subvimageI) (auto intro!: svi[THEN subvimageD]) - -lemma subvimage_trans: - assumes fg: "subvimage A f g" - assumes gh: "subvimage A g h" - shows "subvimage A f h" - by (rule subvimageI) (auto intro!: fg[THEN subvimageD] gh[THEN subvimageD]) - -lemma subvimage_translator: - assumes svi: "subvimage A f g" - shows "\h. \x \ A. h (f x) = g x" -proof (safe intro!: exI[of _ "\x. (THE z. z \ (g ` (f -` {x} \ A)))"]) - fix x assume "x \ A" - show "(THE x'. x' \ (g ` (f -` {f x} \ A))) = g x" - by (rule theI2[of _ "g x"]) - (insert `x \ A`, auto intro!: svi[THEN subvimageD]) -qed - -lemma subvimage_translator_image: - assumes svi: "subvimage A f g" - shows "\h. h ` f ` A = g ` A" -proof - - from subvimage_translator[OF svi] - obtain h where "\x. x \ A \ h (f x) = g x" by auto - thus ?thesis - by (auto intro!: exI[of _ h] - simp: image_compose[symmetric] comp_def cong: image_cong) -qed - -lemma subvimage_finite: - assumes svi: "subvimage A f g" and fin: "finite (f`A)" - shows "finite (g`A)" -proof - - from subvimage_translator_image[OF svi] - obtain h where "g`A = h`f`A" by fastforce - with fin show "finite (g`A)" by simp -qed - -lemma subvimage_disj: - assumes svi: "subvimage A f g" - shows "f -` {x} \ A \ g -` {y} \ A \ - f -` {x} \ g -` {y} \ A = {}" (is "?sub \ ?dist") -proof (rule disjCI) - assume "\ ?dist" - then obtain z where "z \ A" and "x = f z" and "y = g z" by auto - thus "?sub" using svi unfolding subvimage_def by auto -qed - -lemma setsum_image_split: - assumes svi: "subvimage A f g" and fin: "finite (f ` A)" - shows "(\x\f`A. h x) = (\y\g`A. \x\f`(g -` {y} \ A). h x)" - (is "?lhs = ?rhs") -proof - - have "f ` A = - snd ` (SIGMA x : g ` A. f ` (g -` {x} \ A))" - (is "_ = snd ` ?SIGMA") - unfolding image_split_eq_Sigma[symmetric] - by (simp add: image_compose[symmetric] comp_def) - moreover - have snd_inj: "inj_on snd ?SIGMA" - unfolding image_split_eq_Sigma[symmetric] - by (auto intro!: inj_onI subvimageD[OF svi]) - ultimately - have "(\x\f`A. h x) = (\(x,y)\?SIGMA. h y)" - by (auto simp: setsum_reindex intro: setsum_cong) - also have "... = ?rhs" - using subvimage_finite[OF svi fin] fin - apply (subst setsum_Sigma[symmetric]) - by (auto intro!: finite_subset[of _ "f`A"]) - finally show ?thesis . + note XY = simple_distributedI[OF simple_function_Pair[OF X Y] refl] + note YX = simple_distributedI[OF simple_function_Pair[OF Y X] refl] + note simple_distributed_joint_finite[OF this, simp] + let ?f = "\x. prob ((\x. (X x, Y x)) -` {x} \ space M)" + let ?g = "\x. prob ((\x. (Y x, X x)) -` {x} \ space M)" + let ?h = "\x. if x \ (\x. (Y x, X x)) ` space M then prob ((\x. (Y x, X x)) -` {x} \ space M) else 0" + have "\(\x. (X x, Y x)) = - (\x\(\x. (X x, Y x)) ` space M. ?f x * log b (?f x))" + using XY by (rule entropy_simple_distributed) + also have "\ = - (\x\(\(x, y). (y, x)) ` (\x. (X x, Y x)) ` space M. ?g x * log b (?g x))" + by (subst (2) setsum_reindex) (auto simp: inj_on_def intro!: setsum_cong arg_cong[where f="\A. prob A * log b (prob A)"]) + also have "\ = - (\x\(\x. (Y x, X x)) ` space M. ?h x * log b (?h x))" + by (auto intro!: setsum_cong) + also have "\ = entropy b (count_space (Y ` space M) \\<^isub>M count_space (X ` space M)) (\x. (Y x, X x))" + by (subst entropy_distr[OF _ simple_distributed_joint[OF YX]]) + (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite + cong del: setsum_cong intro!: setsum_mono_zero_left) + finally have "\(\x. (X x, Y x)) = entropy b (count_space (Y ` space M) \\<^isub>M count_space (X ` space M)) (\x. (Y x, X x))" . + then show ?thesis + unfolding conditional_entropy_def by simp qed lemma (in information_space) entropy_partition: - assumes sf: "simple_function M X" "simple_function M P" - assumes svi: "subvimage (space M) X P" - shows "\(X) = \(P) + \(X|P)" + assumes X: "simple_function M X" + shows "\(X) = \(f \ X) + \(X|f \ X)" proof - - let ?XP = "\x p. joint_distribution X P {(x, p)}" - let ?X = "\x. distribution X {x}" - let ?P = "\p. distribution P {p}" - note fX = sf(1)[THEN simple_function_imp_finite_random_variable] - note fP = sf(2)[THEN simple_function_imp_finite_random_variable] - note finite_distribution_order[OF fX fP, simp] - have "(\x\X ` space M. ?X x * log b (?X x)) = - (\y\P `space M. \x\X ` space M. ?XP x y * log b (?XP x y))" - proof (subst setsum_image_split[OF svi], - safe intro!: setsum_mono_zero_cong_left imageI) - show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)" - using sf unfolding simple_function_def by auto - next - fix p x assume in_space: "p \ space M" "x \ space M" - assume "?XP (X x) (P p) * log b (?XP (X x) (P p)) \ 0" - hence "(\x. (X x, P x)) -` {(X x, P p)} \ space M \ {}" by (auto simp: distribution_def) - with svi[unfolded subvimage_def, rule_format, OF `x \ space M`] - show "x \ P -` {P p}" by auto - next - fix p x assume in_space: "p \ space M" "x \ space M" - assume "P x = P p" - from this[symmetric] svi[unfolded subvimage_def, rule_format, OF `x \ space M`] - have "X -` {X x} \ space M \ P -` {P p} \ space M" - by auto - hence "(\x. (X x, P x)) -` {(X x, P p)} \ space M = X -` {X x} \ space M" - by auto - thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))" - by (auto simp: distribution_def) - qed - moreover have "\x y. ?XP x y * log b (?XP x y / ?P y) = - ?XP x y * log b (?XP x y) - ?XP x y * log b (?P y)" - by (auto simp add: log_simps zero_less_mult_iff field_simps) - ultimately show ?thesis - unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf] - using setsum_joint_distribution_singleton[OF fX fP] - by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution - setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"]) + note fX = simple_function_compose[OF X, of f] + have eq: "(\x. ((f \ X) x, X x)) ` space M = (\x. (f x, x)) ` X ` space M" by auto + have inj: "\A. inj_on (\x. (f x, x)) A" + by (auto simp: inj_on_def) + show ?thesis + apply (subst entropy_chain_rule[symmetric, OF fX X]) + apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] refl]]) + apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]]) + unfolding eq + apply (subst setsum_reindex[OF inj]) + apply (auto intro!: setsum_cong arg_cong[where f="\A. prob A * log b (prob A)"]) + done qed corollary (in information_space) entropy_data_processing: assumes X: "simple_function M X" shows "\(f \ X) \ \(X)" proof - - note X - 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) + note fX = simple_function_compose[OF X, of f] + from X have "\(X) = \(f\X) + \(X|f\X)" by (rule entropy_partition) then show "\(f \ X) \ \(X)" - by (auto intro: conditional_entropy_positive[OF X fX]) + by (auto intro: conditional_entropy_nonneg[OF X fX]) qed corollary (in information_space) entropy_of_inj: @@ -1411,7 +1313,11 @@ 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]) + unfolding o_assoc + apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]]) + apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\x. prob (X -` {x} \ space M)"]) + apply (auto intro!: setsum_cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def) + done also have "... \ \(f \ X)" using entropy_data_processing[OF sf] . finally show "\(X) \ \(f \ X)" . diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Mon Apr 23 12:14:35 2012 +0200 @@ -6,9 +6,13 @@ header {*Lebesgue Integration*} theory Lebesgue_Integration - imports Measure Borel_Space + imports Measure_Space Borel_Space begin +lemma ereal_minus_eq_PInfty_iff: + fixes x y :: ereal shows "x - y = \ \ y = -\ \ x = \" + by (cases x y rule: ereal2_cases) simp_all + lemma real_ereal_1[simp]: "real (1::ereal) = 1" unfolding one_ereal_def by simp @@ -28,17 +32,17 @@ by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto qed -lemma (in measure_space) measure_Union: +lemma measure_Union: assumes "finite S" "S \ sets M" "\A B. A \ S \ B \ S \ A \ B \ A \ B = {}" - shows "setsum \ S = \ (\S)" + shows "setsum (emeasure M) S = (emeasure M) (\S)" proof - - have "setsum \ S = \ (\i\S. i)" - using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def) - also have "\ = \ (\S)" by (auto intro!: arg_cong[where f=\]) + have "setsum (emeasure M) S = (emeasure M) (\i\S. i)" + using assms by (intro setsum_emeasure[OF _ _ `finite S`]) (auto simp: disjoint_family_on_def) + also have "\ = (emeasure M) (\S)" by (auto intro!: arg_cong[where f="emeasure M"]) finally show ?thesis . qed -lemma (in sigma_algebra) measurable_sets2[intro]: +lemma measurable_sets2[intro]: assumes "f \ measurable M M'" "g \ measurable M M''" and "A \ sets M'" "B \ sets M''" shows "f -` A \ g -` B \ space M \ sets M" @@ -55,7 +59,7 @@ lemma borel_measurable_real_floor: "(\x::real. real \x\) \ borel_measurable borel" - unfolding borel.borel_measurable_iff_ge + unfolding borel_measurable_iff_ge proof (intro allI) fix a :: real { fix x have "a \ real \x\ \ real \a\ \ x" @@ -65,19 +69,7 @@ then show "{w::real \ space borel. a \ real \w\} \ sets borel" by auto qed -lemma measure_preservingD2: - "f \ measure_preserving A B \ f \ measurable A B" - unfolding measure_preserving_def by auto - -lemma measure_preservingD3: - "f \ measure_preserving A B \ f \ space A \ space B" - unfolding measure_preserving_def measurable_def by auto - -lemma measure_preservingD: - "T \ measure_preserving A B \ X \ sets B \ measure A (T -` X \ space A) = measure B X" - unfolding measure_preserving_def by auto - -lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]: +lemma borel_measurable_real_natfloor[intro, simp]: assumes "f \ borel_measurable M" shows "(\x. real (natfloor (f x))) \ borel_measurable M" proof - @@ -87,8 +79,8 @@ show ?thesis by (simp add: comp_def) qed -lemma (in measure_space) AE_not_in: - assumes N: "N \ null_sets" shows "AE x. x \ N" +lemma AE_not_in: + assumes N: "N \ null_sets M" shows "AE x in M. x \ N" using N by (rule AE_I') auto lemma sums_If_finite: @@ -128,7 +120,7 @@ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" -lemma (in sigma_algebra) simple_functionD: +lemma simple_functionD: assumes "simple_function M g" shows "finite (g ` space M)" and "g -` X \ space M \ sets M" proof - @@ -140,7 +132,7 @@ by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def) qed -lemma (in sigma_algebra) simple_function_measurable2[intro]: +lemma simple_function_measurable2[intro]: assumes "simple_function M f" "simple_function M g" shows "f -` A \ g -` B \ space M \ sets M" proof - @@ -149,7 +141,7 @@ then show ?thesis using assms[THEN simple_functionD(2)] by auto qed -lemma (in sigma_algebra) simple_function_indicator_representation: +lemma simple_function_indicator_representation: fixes f ::"'a \ ereal" 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)" @@ -164,7 +156,7 @@ finally show ?thesis by auto qed -lemma (in measure_space) simple_function_notspace: +lemma simple_function_notspace: "simple_function M (\x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") proof - have "?h ` space M \ {0}" unfolding indicator_def by auto @@ -173,7 +165,7 @@ thus ?thesis unfolding simple_function_def by auto qed -lemma (in sigma_algebra) simple_function_cong: +lemma simple_function_cong: assumes "\t. t \ space M \ f t = g t" shows "simple_function M f \ simple_function M g" proof - @@ -183,12 +175,12 @@ thus ?thesis unfolding simple_function_def using assms by simp qed -lemma (in sigma_algebra) simple_function_cong_algebra: +lemma 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: +lemma borel_measurable_simple_function: assumes "simple_function M f" shows "f \ borel_measurable M" proof (rule borel_measurableI) @@ -204,24 +196,23 @@ thus "f -` S \ space M \ sets M" unfolding * . qed -lemma (in sigma_algebra) simple_function_borel_measurable: +lemma simple_function_borel_measurable: fixes f :: "'a \ 'x::{t2_space}" assumes "f \ borel_measurable M" and "finite (f ` space M)" shows "simple_function M f" using assms unfolding simple_function_def by (auto intro: borel_measurable_vimage) -lemma (in sigma_algebra) simple_function_eq_borel_measurable: +lemma simple_function_eq_borel_measurable: fixes f :: "'a \ ereal" 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] + using simple_function_borel_measurable[of f] borel_measurable_simple_function[of M f] by (fastforce simp: simple_function_def) -lemma (in sigma_algebra) simple_function_const[intro, simp]: +lemma simple_function_const[intro, simp]: "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) -lemma (in sigma_algebra) simple_function_compose[intro, simp]: +lemma simple_function_compose[intro, simp]: assumes "simple_function M f" shows "simple_function M (g \ f)" unfolding simple_function_def @@ -238,7 +229,7 @@ by (rule_tac finite_UN) (auto intro!: finite_UN) qed -lemma (in sigma_algebra) simple_function_indicator[intro, simp]: +lemma simple_function_indicator[intro, simp]: assumes "A \ sets M" shows "simple_function M (indicator A)" proof - @@ -250,7 +241,7 @@ using assms by (auto simp: indicator_def [abs_def]) qed -lemma (in sigma_algebra) simple_function_Pair[intro, simp]: +lemma simple_function_Pair[intro, simp]: assumes "simple_function M f" assumes "simple_function M g" shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") @@ -268,13 +259,13 @@ using assms unfolding simple_function_def by auto qed -lemma (in sigma_algebra) simple_function_compose1: +lemma simple_function_compose1: 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: +lemma simple_function_compose2: assumes "simple_function M f" and "simple_function M g" shows "simple_function M (\x. h (f x) (g x))" proof - @@ -283,7 +274,7 @@ thus ?thesis by (simp_all add: comp_def) qed -lemmas (in sigma_algebra) simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] +lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] @@ -291,24 +282,24 @@ and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] and simple_function_max[intro, simp] = simple_function_compose2[where h=max] -lemma (in sigma_algebra) simple_function_setsum[intro, simp]: +lemma simple_function_setsum[intro, simp]: 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) +lemma fixes f g :: "'a \ real" assumes sf: "simple_function M f" shows simple_function_ereal[intro, simp]: "simple_function M (\x. ereal (f x))" by (auto intro!: simple_function_compose1[OF sf]) -lemma (in sigma_algebra) +lemma fixes f g :: "'a \ nat" assumes sf: "simple_function M f" shows simple_function_real_of_nat[intro, simp]: "simple_function M (\x. real (f x))" by (auto intro!: simple_function_compose1[OF sf]) -lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: +lemma borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ ereal" assumes u: "u \ borel_measurable M" shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ @@ -416,14 +407,14 @@ qed (auto simp: divide_nonneg_pos) qed -lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': +lemma borel_measurable_implies_simple_function_sequence': fixes u :: "'a \ ereal" assumes u: "u \ borel_measurable M" obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" using borel_measurable_implies_simple_function_sequence[OF u] by auto -lemma (in sigma_algebra) simple_function_If_set: +lemma simple_function_If_set: assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") proof - @@ -445,7 +436,7 @@ qed qed -lemma (in sigma_algebra) simple_function_If: +lemma simple_function_If: assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" shows "simple_function M (\x. if P x then f x else g x)" proof - @@ -453,58 +444,17 @@ with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp qed -lemma (in measure_space) simple_function_restricted: - fixes f :: "'a \ ereal" assumes "A \ sets M" - 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 f: "finite (f`A) \ finite (?f`space M)" - proof cases - assume "A = space M" - then have "f`A = ?f`space M" by (fastforce simp: image_iff) - then show ?thesis by simp - next - assume "A \ space M" - then obtain x where x: "x \ space M" "x \ A" - using sets_into_space `A \ sets M` by auto - have *: "?f`space M = f`A \ {0}" - proof (auto simp add: image_iff) - show "\x\space M. f x = 0 \ indicator A x = 0" - using x by (auto intro!: bexI[of _ x]) - next - fix x assume "x \ A" - then show "\y\space M. f x = f y * indicator A y" - using `A \ sets M` sets_into_space by (auto intro!: bexI[of _ x]) - next - fix x - assume "indicator A x \ (0::ereal)" - then have "x \ A" by (auto simp: indicator_def split: split_if_asm) - moreover assume "x \ space M" "\y\A. ?f x \ f y" - ultimately show "f x = 0" by auto - qed - then show ?thesis by auto - qed - then show ?thesis - unfolding simple_function_eq_borel_measurable - R.simple_function_eq_borel_measurable - unfolding borel_measurable_restricted[OF `A \ sets M`] - using assms(1)[THEN sets_into_space] - by (auto simp: indicator_def) -qed - -lemma (in sigma_algebra) simple_function_subalgebra: +lemma simple_function_subalgebra: 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'" +lemma simple_function_comp: + assumes T: "T \ measurable M M'" 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)" @@ -523,16 +473,16 @@ section "Simple integral" -definition simple_integral_def: - "integral\<^isup>S M f = (\x \ f ` space M. x * measure M (f -` {x} \ space M))" +definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^isup>S") where + "integral\<^isup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" syntax - "_simple_integral" :: "pttrn \ ereal \ ('a, 'b) measure_space_scheme \ ereal" ("\\<^isup>S _. _ \_" [60,61] 110) + "_simple_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^isup>S _. _ \_" [60,61] 110) translations - "\\<^isup>S x. f \M" == "CONST integral\<^isup>S M (%x. f)" + "\\<^isup>S x. f \M" == "CONST simple_integral M (%x. f)" -lemma (in measure_space) simple_integral_cong: +lemma simple_integral_cong: assumes "\t. t \ space M \ f t = g t" shows "integral\<^isup>S M f = integral\<^isup>S M g" proof - @@ -542,19 +492,8 @@ thus ?thesis unfolding simple_integral_def by simp qed -lemma (in measure_space) simple_integral_cong_measure: - 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 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 \M) = c * \ (space M)" +lemma simple_integral_const[simp]: + "(\\<^isup>Sx. c \M) = c * (emeasure M) (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next @@ -562,9 +501,9 @@ thus ?thesis unfolding simple_integral_def by simp qed -lemma (in measure_space) simple_function_partition: +lemma simple_function_partition: assumes f: "simple_function M f" and g: "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)" + shows "integral\<^isup>S M f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * (emeasure M) A)" (is "_ = setsum _ (?p ` space M)") proof- let ?sub = "\x. ?p ` (f -` {x} \ space M)" @@ -586,13 +525,13 @@ { fix x assume "x \ space M" have "\(?sub (f x)) = (f -` {f x} \ space M)" by auto - with sets have "\ (f -` {f x} \ space M) = setsum \ (?sub (f x))" + with sets have "(emeasure M) (f -` {f x} \ space M) = setsum (emeasure M) (?sub (f x))" by (subst measure_Union) auto } - hence "integral\<^isup>S M f = (\(x,A)\?SIGMA. x * \ A)" + hence "integral\<^isup>S M f = (\(x,A)\?SIGMA. x * (emeasure M) A)" unfolding simple_integral_def using f sets by (subst setsum_Sigma[symmetric]) (auto intro!: setsum_cong setsum_ereal_right_distrib) - also have "\ = (\A\?p ` space M. the_elem (f`A) * \ A)" + also have "\ = (\A\?p ` space M. the_elem (f`A) * (emeasure M) A)" proof - have [simp]: "\x. x \ space M \ f ` ?p x = {f x}" by (auto intro!: imageI) have "(\A. (the_elem (f ` A), A)) ` ?p ` space M @@ -610,7 +549,7 @@ finally show ?thesis . qed -lemma (in measure_space) simple_integral_add[simp]: +lemma simple_integral_add[simp]: assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" shows "(\\<^isup>Sx. f x + g x \M) = integral\<^isup>S M f + integral\<^isup>S M g" proof - @@ -628,7 +567,7 @@ (auto simp add: ereal_left_distrib setsum_addf[symmetric] intro!: setsum_cong) qed -lemma (in measure_space) simple_integral_setsum[simp]: +lemma simple_integral_setsum[simp]: assumes "\i x. i \ P \ 0 \ f i x" 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))" @@ -638,11 +577,11 @@ by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) qed auto -lemma (in measure_space) simple_integral_mult[simp]: +lemma simple_integral_mult[simp]: assumes f: "simple_function M f" "\x. 0 \ f x" 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] f(1)] + note mult = simple_function_mult[OF simple_function_const[of _ c] f(1)] { fix x let ?S = "f -` {f x} \ (\x. c * f x) -` {c * f x} \ space M" assume "x \ space M" hence "(\x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}" @@ -654,9 +593,9 @@ (auto intro!: ereal_0_le_mult setsum_cong simp: mult_assoc) qed -lemma (in measure_space) simple_integral_mono_AE: +lemma simple_integral_mono_AE: assumes f: "simple_function M f" and g: "simple_function M g" - and mono: "AE x. f x \ g x" + and mono: "AE x in M. f x \ g x" 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)" @@ -669,55 +608,55 @@ 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 - show "the_elem (f`?S x) * \ (?S x) \ the_elem (g`?S x) * \ (?S x)" + show "the_elem (f`?S x) * (emeasure M) (?S x) \ the_elem (g`?S x) * (emeasure M) (?S x)" proof (cases "f x \ g x") case True then show ?thesis using * assms(1,2)[THEN simple_functionD(2)] by (auto intro!: ereal_mult_right_mono) next case False - obtain N where N: "{x\space M. \ f x \ g x} \ N" "N \ sets M" "\ N = 0" + obtain N where N: "{x\space M. \ f x \ g x} \ N" "N \ sets M" "(emeasure M) N = 0" using mono by (auto elim!: AE_E) have "?S x \ N" using N `x \ space M` False by auto moreover have "?S x \ sets M" using assms by (rule_tac Int) (auto intro!: simple_functionD) - ultimately have "\ (?S x) \ \ N" - using `N \ sets M` by (auto intro!: measure_mono) - moreover have "0 \ \ (?S x)" + ultimately have "(emeasure M) (?S x) \ (emeasure M) N" + using `N \ sets M` by (auto intro!: emeasure_mono) + moreover have "0 \ (emeasure M) (?S x)" using assms(1,2)[THEN simple_functionD(2)] by auto - ultimately have "\ (?S x) = 0" using `\ N = 0` by auto + ultimately have "(emeasure M) (?S x) = 0" using `(emeasure M) N = 0` by auto then show ?thesis by simp qed qed qed -lemma (in measure_space) simple_integral_mono: +lemma simple_integral_mono: assumes "simple_function M f" and "simple_function M g" and mono: "\ x. x \ space M \ f x \ g x" shows "integral\<^isup>S M f \ integral\<^isup>S M g" using assms by (intro simple_integral_mono_AE) auto -lemma (in measure_space) simple_integral_cong_AE: +lemma simple_integral_cong_AE: assumes "simple_function M f" and "simple_function M g" - and "AE x. f x = g x" + and "AE x in M. 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': +lemma simple_integral_cong': assumes sf: "simple_function M f" "simple_function M g" - and mea: "\ {x\space M. f x \ g x} = 0" + and mea: "(emeasure M) {x\space M. f x \ g x} = 0" 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 "(emeasure M) {x\space M. f x \ g x} = 0" by fact show "{x \ space M. f x \ g x} \ sets M" using sf[THEN borel_measurable_simple_function] by auto qed simp -lemma (in measure_space) simple_integral_indicator: +lemma simple_integral_indicator: assumes "A \ sets M" assumes "simple_function M f" shows "(\\<^isup>Sx. f x * indicator A x \M) = - (\x \ f ` space M. x * \ (f -` {x} \ space M \ A))" + (\x \ f ` space M. x * (emeasure M) (f -` {x} \ space M \ A))" proof cases assume "A = space M" moreover hence "(\\<^isup>Sx. f x * indicator A x \M) = integral\<^isup>S M f" @@ -737,7 +676,7 @@ show "0 \ ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x]) qed have *: "(\\<^isup>Sx. f x * indicator A x \M) = - (\x \ f ` space M \ {0}. x * \ (f -` {x} \ space M \ A))" + (\x \ f ` space M \ {0}. x * (emeasure M) (f -` {x} \ space M \ A))" unfolding simple_integral_def I proof (rule setsum_mono_zero_cong_left) show "finite (f ` space M \ {0})" @@ -747,118 +686,83 @@ have "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" by (auto simp: image_iff) thus "\i\f ` space M \ {0} - (f ` A \ {0}). - i * \ (f -` {i} \ space M \ A) = 0" by auto + i * (emeasure M) (f -` {i} \ space M \ A) = 0" by auto next fix x assume "x \ f`A \ {0}" hence "x \ 0 \ ?I -` {x} \ space M = f -` {x} \ space M \ A" by (auto simp: indicator_def split: split_if_asm) - thus "x * \ (?I -` {x} \ space M) = - x * \ (f -` {x} \ space M \ A)" by (cases "x = 0") simp_all + thus "x * (emeasure M) (?I -` {x} \ space M) = + x * (emeasure M) (f -` {x} \ space M \ A)" by (cases "x = 0") simp_all qed show ?thesis unfolding * using assms(2) unfolding simple_function_def by (auto intro!: setsum_mono_zero_cong_right) qed -lemma (in measure_space) simple_integral_indicator_only[simp]: +lemma simple_integral_indicator_only[simp]: assumes "A \ sets M" - shows "integral\<^isup>S M (indicator A) = \ A" + shows "integral\<^isup>S M (indicator A) = emeasure M 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 next assume "space M \ {}" hence "(\x. 1) ` space M = {1::ereal}" by auto thus ?thesis - using simple_integral_indicator[OF assms simple_function_const[of 1]] + using simple_integral_indicator[OF assms simple_function_const[of _ 1]] using sets_into_space[OF assms] - by (auto intro!: arg_cong[where f="\"]) + by (auto intro!: arg_cong[where f="(emeasure M)"]) qed -lemma (in measure_space) simple_integral_null_set: - assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets" +lemma simple_integral_null_set: + assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" shows "(\\<^isup>Sx. u x * indicator N x \M) = 0" proof - - have "AE x. indicator N x = (0 :: ereal)" - using `N \ null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N]) + have "AE x in M. indicator N x = (0 :: ereal)" + using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) then have "(\\<^isup>Sx. u x * indicator N x \M) = (\\<^isup>Sx. 0 \M)" using assms apply (intro simple_integral_cong_AE) by auto then show ?thesis by simp qed -lemma (in measure_space) simple_integral_cong_AE_mult_indicator: - assumes sf: "simple_function M f" and eq: "AE x. x \ S" and "S \ sets M" +lemma simple_integral_cong_AE_mult_indicator: + assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" shows "integral\<^isup>S M f = (\\<^isup>Sx. f x * indicator S x \M)" using assms by (intro simple_integral_cong_AE) auto -lemma (in measure_space) simple_integral_restricted: - assumes "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") +lemma simple_integral_distr: + assumes T: "T \ measurable M M'" + and f: "simple_function M' f" + shows "integral\<^isup>S (distr M M' T) f = (\\<^isup>S x. f (T x) \M)" unfolding simple_integral_def -proof (simp, safe intro!: setsum_mono_zero_cong_left) - from sf show "finite (?f ` space M)" - unfolding simple_function_def by auto +proof (intro setsum_mono_zero_cong_right ballI) + show "(\x. f (T x)) ` space M \ f ` space (distr M M' T)" + using T unfolding measurable_def by auto + show "finite (f ` space (distr M M' T))" + using f unfolding simple_function_def by auto next - fix x assume "x \ A" - then show "f x \ ?f ` space M" - using sets_into_space `A \ sets M` by (auto intro!: image_eqI[of _ _ x]) -next - fix x assume "x \ space M" "?f x \ f`A" - then have "x \ A" by (auto simp: image_iff) - then show "?f x * \ (?f -` {?f x} \ space M) = 0" by simp + fix i assume "i \ f ` space (distr M M' T) - (\x. f (T x)) ` space M" + then have "T -` (f -` {i} \ space (distr M M' T)) \ space M = {}" by (auto simp: image_iff) + with f[THEN simple_functionD(2), of "{i}"] + show "i * emeasure (distr M M' T) (f -` {i} \ space (distr M M' T)) = 0" + using T by (simp add: emeasure_distr) next - fix x assume "x \ A" - then have "f x \ 0 \ - f -` {f x} \ A = ?f -` {f x} \ space M" - using `A \ sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show "f x * \ (f -` {f x} \ A) = - f x * \ (?f -` {f x} \ space M)" - unfolding ereal_mult_cancel_left by auto + 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 + with f[THEN simple_functionD(2), of "{i}"] T + show "i * emeasure (distr M M' T) (f -` {i} \ space (distr M M' T)) = + i * (emeasure M) ((\x. f (T x)) -` {i} \ space M)" + by (auto simp add: emeasure_distr) qed -lemma (in measure_space) simple_integral_subalgebra: - 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 [abs_def] by simp - -lemma (in measure_space) simple_integral_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - and f: "simple_function M' f" - shows "integral\<^isup>S M' f = (\\<^isup>S x. f (T x) \M)" -proof - - interpret T: measure_space M' 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 measure_preserving_def by auto - show "finite (f ` space M')" - 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) - with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{i}"] - 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 measure_preserving_def by auto - with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{i}"] - show "i * T.\ (f -` {i} \ space M') = i * \ ((\x. f (T x)) -` {i} \ space M)" - by auto - qed -qed - -lemma (in measure_space) simple_integral_cmult_indicator: +lemma simple_integral_cmult_indicator: assumes A: "A \ sets M" - shows "(\\<^isup>Sx. c * indicator A x \M) = c * \ A" + shows "(\\<^isup>Sx. c * indicator A x \M) = c * (emeasure M) A" using simple_integral_mult[OF simple_function_indicator[OF A]] unfolding simple_integral_indicator_only[OF A] by simp -lemma (in measure_space) simple_integral_positive: - assumes f: "simple_function M f" and ae: "AE x. 0 \ f x" +lemma simple_integral_positive: + assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" shows "0 \ integral\<^isup>S M f" proof - have "integral\<^isup>S M (\x. 0) \ integral\<^isup>S M f" @@ -868,29 +772,23 @@ section "Continuous positive integration" -definition positive_integral_def: +definition positive_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^isup>P") where "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^isup>S M g)" syntax - "_positive_integral" :: "pttrn \ ereal \ ('a, 'b) measure_space_scheme \ ereal" ("\\<^isup>+ _. _ \_" [60,61] 110) + "_positive_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^isup>+ _. _ \_" [60,61] 110) translations - "\\<^isup>+ x. f \M" == "CONST integral\<^isup>P M (%x. f)" + "\\<^isup>+ x. f \M" == "CONST positive_integral M (%x. f)" -lemma (in measure_space) positive_integral_cong_measure: - 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" - unfolding positive_integral_def - unfolding simple_function_cong_algebra[OF assms(2,3), symmetric] - using AE_cong_measure[OF assms] - using simple_integral_cong_measure[OF assms] - by (auto intro!: SUP_cong) - -lemma (in measure_space) positive_integral_positive: +lemma positive_integral_positive: "0 \ integral\<^isup>P M f" by (auto intro!: SUP_upper2[of "\x. 0"] simp: positive_integral_def le_fun_def) -lemma (in measure_space) positive_integral_def_finite: +lemma positive_integral_not_MInfty[simp]: "integral\<^isup>P M f \ -\" + using positive_integral_positive[of M f] by auto + +lemma positive_integral_def_finite: "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^isup>S M g)" (is "_ = SUPR ?A ?f") unfolding positive_integral_def @@ -898,7 +796,7 @@ fix g assume g: "simple_function M g" "g \ max 0 \ f" let ?G = "{x \ space M. \ g x \ \}" note gM = g(1)[THEN borel_measurable_simple_function] - have \G_pos: "0 \ \ ?G" using gM by auto + have \G_pos: "0 \ (emeasure M) ?G" using gM by auto let ?g = "\y x. if g x = \ then y else max 0 (g x)" from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" apply (safe intro!: simple_function_max simple_function_If) @@ -907,21 +805,22 @@ show "integral\<^isup>S M g \ SUPR ?A ?f" proof cases have g0: "?g 0 \ ?A" by (intro g_in_A) auto - assume "\ ?G = 0" - with gM have "AE x. x \ ?G" by (simp add: AE_iff_null_set) + assume "(emeasure M) ?G = 0" + with gM have "AE x in M. x \ ?G" + by (auto simp add: AE_iff_null intro!: null_setsI) with gM g show ?thesis by (intro SUP_upper2[OF g0] simple_integral_mono_AE) (auto simp: max_def intro!: simple_function_If) next - assume \G: "\ ?G \ 0" + assume \G: "(emeasure M) ?G \ 0" have "SUPR ?A (integral\<^isup>S M) = \" proof (intro SUP_PInfty) fix n :: nat - let ?y = "ereal (real n) / (if \ ?G = \ then 1 else \ ?G)" + let ?y = "ereal (real n) / (if (emeasure M) ?G = \ then 1 else (emeasure M) ?G)" have "0 \ ?y" "?y \ \" using \G \G_pos by (auto simp: ereal_divide_eq) then have "?g ?y \ ?A" by (rule g_in_A) - have "real n \ ?y * \ ?G" - using \G \G_pos by (cases "\ ?G") (auto simp: field_simps) + have "real n \ ?y * (emeasure M) ?G" + using \G \G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) also have "\ = (\\<^isup>Sx. ?y * indicator ?G x \M)" using `0 \ ?y` `?g ?y \ ?A` gM by (subst simple_integral_cmult_indicator) auto @@ -934,15 +833,15 @@ qed qed (auto intro: SUP_upper) -lemma (in measure_space) positive_integral_mono_AE: - assumes ae: "AE x. u x \ v x" shows "integral\<^isup>P M u \ integral\<^isup>P M v" +lemma positive_integral_mono_AE: + assumes ae: "AE x in M. u x \ v x" shows "integral\<^isup>P M u \ integral\<^isup>P M v" unfolding positive_integral_def proof (safe intro!: SUP_mono) fix n assume n: "simple_function M n" "n \ max 0 \ u" from ae[THEN AE_E] guess N . note N = this - then have ae_N: "AE x. x \ N" by (auto intro: AE_not_in) + then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) let ?n = "\x. n x * indicator (space M - N) x" - have "AE x. n x \ ?n x" "simple_function M ?n" + have "AE x in M. n x \ ?n x" "simple_function M ?n" using n N ae_N by auto moreover { fix x have "?n x \ max 0 (v x)" @@ -959,19 +858,19 @@ by force qed -lemma (in measure_space) positive_integral_mono: +lemma positive_integral_mono: "(\x. x \ space M \ u x \ v x) \ integral\<^isup>P M u \ integral\<^isup>P M v" by (auto intro: positive_integral_mono_AE) -lemma (in measure_space) positive_integral_cong_AE: - "AE x. u x = v x \ integral\<^isup>P M u = integral\<^isup>P M v" +lemma positive_integral_cong_AE: + "AE x in M. 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_cong: +lemma positive_integral_cong: "(\x. x \ space M \ u x = v x) \ integral\<^isup>P M u = integral\<^isup>P M v" by (auto intro: positive_integral_cong_AE) -lemma (in measure_space) positive_integral_eq_simple_integral: +lemma positive_integral_eq_simple_integral: assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^isup>P M f = integral\<^isup>S M f" proof - let ?f = "\x. f x * indicator (space M) x" @@ -987,10 +886,10 @@ by (simp cong: positive_integral_cong simple_integral_cong) qed -lemma (in measure_space) positive_integral_eq_simple_integral_AE: - assumes f: "simple_function M f" "AE x. 0 \ f x" shows "integral\<^isup>P M f = integral\<^isup>S M f" +lemma positive_integral_eq_simple_integral_AE: + assumes f: "simple_function M f" "AE x in M. 0 \ f x" shows "integral\<^isup>P M f = integral\<^isup>S M f" proof - - have "AE x. f x = max 0 (f x)" using f by (auto split: split_max) + have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) with f have "integral\<^isup>P M f = integral\<^isup>S M (\x. max 0 (f x))" by (simp cong: positive_integral_cong_AE simple_integral_cong_AE add: positive_integral_eq_simple_integral) @@ -998,7 +897,7 @@ by (auto intro!: simple_integral_cong_AE split: split_max) qed -lemma (in measure_space) positive_integral_SUP_approx: +lemma positive_integral_SUP_approx: assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" shows "integral\<^isup>S M u \ (SUP i. integral\<^isup>P M (f i))" (is "_ \ ?S") @@ -1028,7 +927,7 @@ note B_u = Int[OF u(1)[THEN simple_functionD(2)] B] let ?B' = "\i n. (u -` {i} \ space M) \ ?B n" - have measure_conv: "\i. \ (u -` {i} \ space M) = (SUP n. \ (?B' i n))" + have measure_conv: "\i. (emeasure M) (u -` {i} \ space M) = (SUP n. (emeasure M) (?B' i n))" proof - fix i have 1: "range (?B' i) \ sets M" using B_u by auto @@ -1051,17 +950,17 @@ thus ?thesis using `x \ space M` by auto qed qed - then show "?thesis i" using continuity_from_below[OF 1 2] by simp + then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp qed 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_ereal_setsum, safe) fix x n assume "x \ space M" - with u_range show "incseq (\i. u x * \ (?B' (u x) i))" "\i. 0 \ u x * \ (?B' (u x) i)" - using B_mono B_u by (auto intro!: measure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) + with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" + using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) next - show "integral\<^isup>S M u = (\i\u ` space M. SUP n. i * \ (?B' i n))" + show "integral\<^isup>S M u = (\i\u ` space M. SUP n. i * (emeasure M) (?B' i n))" using measure_conv u_range B_u unfolding simple_integral_def by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric]) qed @@ -1089,7 +988,7 @@ ultimately show "a * integral\<^isup>S M u \ ?S" by simp qed -lemma (in measure_space) incseq_positive_integral: +lemma incseq_positive_integral: assumes "incseq f" shows "incseq (\i. integral\<^isup>P M (f i))" proof - have "\i x. f i x \ f (Suc i) x" @@ -1099,7 +998,7 @@ qed text {* Beppo-Levi monotone convergence theorem *} -lemma (in measure_space) positive_integral_monotone_convergence_SUP: +lemma positive_integral_monotone_convergence_SUP: assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" proof (rule antisym) @@ -1107,7 +1006,7 @@ by (auto intro!: SUP_least SUP_upper positive_integral_mono) next show "(\\<^isup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^isup>P M (f j))" - unfolding positive_integral_def_finite[of "\x. SUP i. f i x"] + unfolding positive_integral_def_finite[of _ "\x. SUP i. f i x"] proof (safe intro!: SUP_least) fix g assume g: "simple_function M g" and "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" @@ -1119,15 +1018,15 @@ qed qed -lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE: - assumes f: "\i. AE x. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" +lemma positive_integral_monotone_convergence_SUP_AE: + assumes f: "\i. AE x in M. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" proof - - from f have "AE x. \i. f i x \ f (Suc i) x \ 0 \ f i x" + from f have "AE x in M. \i. f i x \ f (Suc i) x \ 0 \ f i x" by (simp add: AE_all_countable) from this[THEN AE_E] guess N . note N = this let ?f = "\i x. if x \ space M - N then f i x else 0" - have f_eq: "AE x. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N]) + have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) then have "(\\<^isup>+ x. (SUP i. f i x) \M) = (\\<^isup>+ x. (SUP i. ?f i x) \M)" by (auto intro!: positive_integral_cong_AE) also have "\ = (SUP i. (\\<^isup>+ x. ?f i x \M))" @@ -1143,14 +1042,14 @@ finally show ?thesis . qed -lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq: - assumes f: "incseq f" "\i. AE x. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" +lemma positive_integral_monotone_convergence_SUP_AE_incseq: + assumes f: "incseq f" "\i. AE x in M. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" using f[unfolded incseq_Suc_iff le_fun_def] by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel]) auto -lemma (in measure_space) positive_integral_monotone_convergence_simple: +lemma positive_integral_monotone_convergence_simple: assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" shows "(SUP i. integral\<^isup>S M (f i)) = (\\<^isup>+x. (SUP i. f i x) \M)" using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1) @@ -1161,7 +1060,7 @@ "(\\<^isup>+x. max 0 (f x) \M) = integral\<^isup>P M f" by (simp add: le_fun_def positive_integral_def) -lemma (in measure_space) positive_integral_cong_pos: +lemma positive_integral_cong_pos: assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" shows "integral\<^isup>P M f = integral\<^isup>P M g" proof - @@ -1174,10 +1073,10 @@ then show ?thesis by (simp add: positive_integral_max_0) qed -lemma (in measure_space) SUP_simple_integral_sequences: +lemma SUP_simple_integral_sequences: assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" - and eq: "AE x. (SUP i. f i x) = (SUP i. g i x)" + and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))" (is "SUPR _ ?F = SUPR _ ?G") proof - @@ -1190,32 +1089,11 @@ finally show ?thesis by simp qed -lemma (in measure_space) positive_integral_const[simp]: - "0 \ c \ (\\<^isup>+ x. c \M) = c * \ (space M)" +lemma positive_integral_const[simp]: + "0 \ c \ (\\<^isup>+ x. c \M) = c * (emeasure M) (space M)" by (subst positive_integral_eq_simple_integral) auto -lemma (in measure_space) positive_integral_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - and f: "f \ borel_measurable M'" - shows "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" -proof - - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) - from T.borel_measurable_implies_simple_function_sequence'[OF f] - guess f' . note f' = this - let ?f = "\i x. f' i (T x)" - have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def) - have sup: "\x. (SUP i. ?f i x) = max 0 (f (T x))" - using f'(4) . - have sf: "\i. simple_function M (\x. f' i (T x))" - using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(1)] . - show "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" - using - T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)] - positive_integral_monotone_convergence_simple[OF inc f'(5) sf] - by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f') -qed - -lemma (in measure_space) positive_integral_linear: +lemma positive_integral_linear: assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" and g: "g \ borel_measurable M" "\x. 0 \ g x" shows "(\\<^isup>+ x. a * f x + g x \M) = a * integral\<^isup>P M f + integral\<^isup>P M g" @@ -1254,7 +1132,7 @@ by (subst SUPR_ereal_cmult[symmetric, OF u(6) `0 \ a`]) (auto intro!: SUPR_ereal_add simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg) } - then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)" + then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg) qed @@ -1268,8 +1146,8 @@ then show ?thesis by (simp add: positive_integral_max_0) qed -lemma (in measure_space) positive_integral_cmult: - assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" "0 \ c" +lemma positive_integral_cmult: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ c" shows "(\\<^isup>+ x. c * f x \M) = c * integral\<^isup>P M f" proof - have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` @@ -1277,68 +1155,68 @@ have "(\\<^isup>+ x. c * f x \M) = (\\<^isup>+ x. c * max 0 (f x) \M)" by (simp add: positive_integral_max_0) then show ?thesis - using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" "\x. 0"] f + using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f by (auto simp: positive_integral_max_0) qed -lemma (in measure_space) positive_integral_multc: - assumes "f \ borel_measurable M" "AE x. 0 \ f x" "0 \ c" +lemma positive_integral_multc: + assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ 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\M) = \ A" +lemma positive_integral_indicator[simp]: + "A \ sets M \ (\\<^isup>+ x. indicator A x\M) = (emeasure 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: - "0 \ c \ A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * \ A" +lemma positive_integral_cmult_indicator: + "0 \ c \ A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * (emeasure M) 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: "f \ borel_measurable M" "AE x. 0 \ f x" - and g: "g \ borel_measurable M" "AE x. 0 \ g x" +lemma positive_integral_add: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" shows "(\\<^isup>+ x. f x + g x \M) = integral\<^isup>P M f + integral\<^isup>P M g" proof - - have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" + have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg) have "(\\<^isup>+ x. f x + g x \M) = (\\<^isup>+ x. max 0 (f x + g x) \M)" by (simp add: positive_integral_max_0) also have "\ = (\\<^isup>+ x. max 0 (f x) + max 0 (g x) \M)" unfolding ae[THEN positive_integral_cong_AE] .. also have "\ = (\\<^isup>+ x. max 0 (f x) \M) + (\\<^isup>+ x. max 0 (g x) \M)" - using positive_integral_linear[of "\x. max 0 (f x)" 1 "\x. max 0 (g x)"] f g + using positive_integral_linear[of "\x. max 0 (f x)" _ 1 "\x. max 0 (g x)"] f g by auto finally show ?thesis by (simp add: positive_integral_max_0) qed -lemma (in measure_space) positive_integral_setsum: - assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x. 0 \ f i x" +lemma positive_integral_setsum: + assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x in M. 0 \ f i x" shows "(\\<^isup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^isup>P M (f i))" proof cases assume f: "finite P" - from assms have "AE x. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto + from assms have "AE x in M. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto from f this assms(1) show ?thesis proof induct case (insert i P) - then have "f i \ borel_measurable M" "AE x. 0 \ f i x" - "(\x. \i\P. f i x) \ borel_measurable M" "AE x. 0 \ (\i\P. f i x)" + then have "f i \ borel_measurable M" "AE x in M. 0 \ f i x" + "(\x. \i\P. f i x) \ borel_measurable M" "AE x in M. 0 \ (\i\P. f i x)" by (auto intro!: borel_measurable_ereal_setsum setsum_nonneg) from positive_integral_add[OF this] show ?case using insert by auto qed simp qed simp -lemma (in measure_space) positive_integral_Markov_inequality: - assumes u: "u \ borel_measurable M" "AE x. 0 \ u x" and "A \ sets M" and c: "0 \ c" "c \ \" - shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" - (is "\ ?A \ _ * ?PI") +lemma positive_integral_Markov_inequality: + assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ c" "c \ \" + shows "(emeasure M) ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" + (is "(emeasure M) ?A \ _ * ?PI") proof - have "?A \ sets M" using `A \ sets M` u by auto - hence "\ ?A = (\\<^isup>+ x. indicator ?A x \M)" + hence "(emeasure M) ?A = (\\<^isup>+ x. indicator ?A x \M)" using positive_integral_indicator by simp also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x) \M)" using u c by (auto intro!: positive_integral_mono_AE @@ -1349,17 +1227,17 @@ finally show ?thesis . qed -lemma (in measure_space) positive_integral_noteq_infinite: - assumes g: "g \ borel_measurable M" "AE x. 0 \ g x" +lemma positive_integral_noteq_infinite: + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" and "integral\<^isup>P M g \ \" - shows "AE x. g x \ \" + shows "AE x in M. g x \ \" proof (rule ccontr) - assume c: "\ (AE x. g x \ \)" - have "\ {x\space M. g x = \} \ 0" - using c g by (simp add: AE_iff_null_set) - moreover have "0 \ \ {x\space M. g x = \}" using g by (auto intro: measurable_sets) - ultimately have "0 < \ {x\space M. g x = \}" by auto - then have "\ = \ * \ {x\space M. g x = \}" by auto + assume c: "\ (AE x in M. g x \ \)" + have "(emeasure M) {x\space M. g x = \} \ 0" + using c g by (auto simp add: AE_iff_null) + moreover have "0 \ (emeasure M) {x\space M. g x = \}" using g by (auto intro: measurable_sets) + ultimately have "0 < (emeasure M) {x\space M. g x = \}" by auto + then have "\ = \ * (emeasure M) {x\space M. g x = \}" by auto also have "\ \ (\\<^isup>+x. \ * indicator {x\space M. g x = \} x \M)" using g by (subst positive_integral_cmult_indicator) auto also have "\ \ integral\<^isup>P M g" @@ -1367,34 +1245,34 @@ finally show False using `integral\<^isup>P M g \ \` by auto qed -lemma (in measure_space) positive_integral_diff: +lemma positive_integral_diff: assumes f: "f \ borel_measurable M" - and g: "g \ borel_measurable M" "AE x. 0 \ g x" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" and fin: "integral\<^isup>P M g \ \" - and mono: "AE x. g x \ f x" + and mono: "AE x in M. g x \ f x" shows "(\\<^isup>+ x. f x - g x \M) = integral\<^isup>P M f - integral\<^isup>P M g" proof - - have diff: "(\x. f x - g x) \ borel_measurable M" "AE x. 0 \ f x - g x" + have diff: "(\x. f x - g x) \ borel_measurable M" "AE x in M. 0 \ f x - g x" using assms by (auto intro: ereal_diff_positive) - have pos_f: "AE x. 0 \ f x" using mono g by auto + have pos_f: "AE x in M. 0 \ f x" using mono g by auto { fix a b :: ereal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" by (cases rule: ereal2_cases[of a b]) auto } note * = this - then have "AE x. f x = f x - g x + g x" + then have "AE x in M. f x = f x - g x + g x" using mono positive_integral_noteq_infinite[OF g fin] assms by auto then have **: "integral\<^isup>P M f = (\\<^isup>+x. f x - g x \M) + integral\<^isup>P M g" unfolding positive_integral_add[OF diff g, symmetric] by (rule positive_integral_cong_AE) show ?thesis unfolding ** - using fin positive_integral_positive[of g] + using fin positive_integral_positive[of M g] by (cases rule: ereal2_cases[of "\\<^isup>+ x. f x - g x \M" "integral\<^isup>P M g"]) auto qed -lemma (in measure_space) positive_integral_suminf: - assumes f: "\i. f i \ borel_measurable M" "\i. AE x. 0 \ f i x" +lemma positive_integral_suminf: + assumes f: "\i. f i \ borel_measurable M" "\i. AE x in M. 0 \ f i x" shows "(\\<^isup>+ x. (\i. f i x) \M) = (\i. integral\<^isup>P M (f i))" proof - - have all_pos: "AE x. \i. 0 \ f i x" + have all_pos: "AE x in M. \i. 0 \ f i x" using assms by (auto simp: AE_all_countable) have "(\i. integral\<^isup>P M (f i)) = (SUP n. \iP M (f i))" using positive_integral_positive by (rule suminf_ereal_eq_SUPR) @@ -1409,12 +1287,12 @@ qed text {* Fatou's lemma: convergence theorem on limes inferior *} -lemma (in measure_space) positive_integral_lim_INF: +lemma positive_integral_lim_INF: fixes u :: "nat \ 'a \ ereal" - assumes u: "\i. u i \ borel_measurable M" "\i. AE x. 0 \ u i x" + assumes u: "\i. u i \ borel_measurable M" "\i. AE x in M. 0 \ u i x" shows "(\\<^isup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^isup>P M (u n))" proof - - have pos: "AE x. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) + have pos: "AE x in M. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) have "(\\<^isup>+ x. liminf (\n. u n x) \M) = (SUP n. \\<^isup>+ x. (INF i:{n..}. u i x) \M)" unfolding liminf_SUPR_INFI using pos u @@ -1426,120 +1304,31 @@ finally show ?thesis . qed -lemma (in measure_space) measure_space_density: - assumes u: "u \ borel_measurable M" "AE x. 0 \ u x" - 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 - have pos: "\A. AE x. 0 \ u x * indicator A x" - using u by (auto simp: ereal_zero_le_0_iff) - then show "positive M' (measure M')" unfolding M' - using u(1) by (auto simp: positive_def intro!: positive_integral_positive) - 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 u by (auto intro: borel_measurable_indicator) - assume disj: "disjoint_family A" - have "(\n. measure M' (A n)) = (\\<^isup>+ x. (\n. u x * indicator (A n) x) \M)" - unfolding M' using u(1) * - by (simp add: positive_integral_suminf[OF _ pos, symmetric]) - also have "\ = (\\<^isup>+ x. u x * (\n. indicator (A n) x) \M)" using u - by (intro positive_integral_cong_AE) - (elim AE_mp, auto intro!: AE_I2 suminf_cmult_ereal) - also have "\ = (\\<^isup>+ x. u x * indicator (\n. A n) x \M)" - unfolding suminf_indicator[OF disj] .. - finally show "(\n. measure M' (A n)) = measure M' (\x. A x)" - unfolding M' by simp - qed - qed -qed - -lemma (in measure_space) positive_integral_null_set: - assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x \M) = 0" +lemma positive_integral_null_set: + assumes "N \ null_sets M" shows "(\\<^isup>+ x. u x * indicator N x \M) = 0" proof - 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) - show "\ N = 0" "N \ sets M" + show "(emeasure M) N = 0" "N \ sets M" using assms by auto qed then show ?thesis by simp qed -lemma (in measure_space) positive_integral_translated_density: - assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" - assumes g: "g \ borel_measurable M" "AE x. 0 \ g x" - 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 f 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 [abs_def] by (auto simp: M') - from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this - note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)] - note G'(2)[simp] - { fix P have "AE x. P x \ AE x in M'. P x" - using positive_integral_null_set[of _ f] - unfolding T.almost_everywhere_def almost_everywhere_def - by (auto simp: M') } - note ac = this - from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x" - by (auto intro!: ac split: split_max) - { fix i - let ?I = "\y x. indicator (G i -` {y} \ space M) x" - { fix x assume *: "x \ space M" "0 \ f x" "0 \ g x" - then have [simp]: "G i ` space M \ {y. G i x = y \ x \ space M} = {G i x}" by auto - from * G' G have "(\y\G i`space M. y * (f x * ?I y x)) = f x * (\y\G i`space M. (y * ?I y x))" - by (subst setsum_ereal_right_distrib) (auto simp: ac_simps) - also have "\ = f x * G i x" - by (simp add: indicator_def if_distrib setsum_cases) - finally have "(\y\G i`space M. y * (f x * ?I y x)) = f x * G i x" . } - note to_singleton = this - 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 "\ = (\y\G i`space M. y * (\\<^isup>+x. f x * ?I y x \M))" - unfolding simple_integral_def M' by simp - also have "\ = (\y\G i`space M. (\\<^isup>+x. y * (f x * ?I y x) \M))" - using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric]) - also have "\ = (\\<^isup>+x. (\y\G i`space M. y * (f x * ?I y x)) \M)" - using f G' G by (auto intro!: positive_integral_setsum[symmetric]) - finally have "integral\<^isup>P M' (G i) = (\\<^isup>+x. f x * G i x \M)" - using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) } - note [simp] = this - have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G - using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`] - by (simp cong: T.positive_integral_cong_AE) - also have "\ = (SUP i. (\\<^isup>+x. f x * G i x \M))" by simp - also have "\ = (\\<^isup>+x. (SUP i. f x * G i x) \M)" - using f G' G(2)[THEN incseq_SucD] G - by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) - (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff) - also have "\ = (\\<^isup>+x. f x * g x \M)" using f G' G g - by (intro positive_integral_cong_AE) - (auto simp add: SUPR_ereal_cmult split: split_max) - finally show "integral\<^isup>P M' g = (\\<^isup>+x. f x * g x \M)" . -qed - -lemma (in measure_space) positive_integral_0_iff: - assumes u: "u \ borel_measurable M" and pos: "AE x. 0 \ u x" - shows "integral\<^isup>P M u = 0 \ \ {x\space M. u x \ 0} = 0" - (is "_ \ \ ?A = 0") +lemma positive_integral_0_iff: + assumes u: "u \ borel_measurable M" and pos: "AE x in M. 0 \ u x" + shows "integral\<^isup>P M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" + (is "_ \ (emeasure M) ?A = 0") proof - have u_eq: "(\\<^isup>+ x. u x * indicator ?A x \M) = integral\<^isup>P M u" by (auto intro!: positive_integral_cong simp: indicator_def) show ?thesis proof - assume "\ ?A = 0" - with positive_integral_null_set[of ?A u] u - show "integral\<^isup>P M u = 0" by (simp add: u_eq) + assume "(emeasure M) ?A = 0" + with positive_integral_null_set[of ?A M u] u + show "integral\<^isup>P M u = 0" by (simp add: u_eq null_sets_def) next { fix r :: ereal and n :: nat assume gt_1: "1 \ real n * r" then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) @@ -1547,17 +1336,17 @@ note gt_1 = this assume *: "integral\<^isup>P M u = 0" let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" - have "0 = (SUP n. \ (?M n \ ?A))" + have "0 = (SUP n. (emeasure M) (?M n \ ?A))" proof - { fix n :: nat from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] - have "\ (?M n \ ?A) \ 0" unfolding u_eq * using u by simp - moreover have "0 \ \ (?M n \ ?A)" using u by auto - ultimately have "\ (?M n \ ?A) = 0" by auto } + have "(emeasure M) (?M n \ ?A) \ 0" unfolding u_eq * using u by simp + moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto + ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } thus ?thesis by simp qed - also have "\ = \ (\n. ?M n \ ?A)" - proof (safe intro!: continuity_from_below) + also have "\ = (emeasure M) (\n. ?M n \ ?A)" + proof (safe intro!: SUP_emeasure_incseq) fix n show "?M n \ ?A \ sets M" using u by (auto intro!: Int) next @@ -1570,8 +1359,8 @@ finally show "1 \ real (Suc n) * u x" by auto qed qed - also have "\ = \ {x\space M. 0 < u x}" - proof (safe intro!: arg_cong[where f="\"] dest!: gt_1) + also have "\ = (emeasure M) {x\space M. 0 < u x}" + proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) fix x assume "0 < u x" and [simp, intro]: "x \ space M" show "x \ (\n. ?M n \ ?A)" proof (cases "u x") @@ -1582,88 +1371,48 @@ thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) qed (insert `0 < u x`, auto) qed auto - finally have "\ {x\space M. 0 < u x} = 0" by simp + finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp moreover - from pos have "AE x. \ (u x < 0)" by auto - then have "\ {x\space M. u x < 0} = 0" - using AE_iff_null_set u by auto - moreover have "\ {x\space M. u x \ 0} = \ {x\space M. u x < 0} + \ {x\space M. 0 < u x}" - using u by (subst measure_additive) (auto intro!: arg_cong[where f=\]) - ultimately show "\ ?A = 0" by simp + from pos have "AE x in M. \ (u x < 0)" by auto + then have "(emeasure M) {x\space M. u x < 0} = 0" + using AE_iff_null[of M] u by auto + moreover have "(emeasure M) {x\space M. u x \ 0} = (emeasure M) {x\space M. u x < 0} + (emeasure M) {x\space M. 0 < u x}" + using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) + ultimately show "(emeasure M) ?A = 0" by simp qed qed -lemma (in measure_space) positive_integral_0_iff_AE: +lemma positive_integral_0_iff_AE: assumes u: "u \ borel_measurable M" - shows "integral\<^isup>P M u = 0 \ (AE x. u x \ 0)" + shows "integral\<^isup>P M u = 0 \ (AE x in M. u x \ 0)" proof - have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" using u by auto from positive_integral_0_iff[of "\x. max 0 (u x)"] - have "integral\<^isup>P M u = 0 \ (AE x. max 0 (u x) = 0)" + have "integral\<^isup>P M u = 0 \ (AE x in M. max 0 (u x) = 0)" unfolding positive_integral_max_0 - using AE_iff_null_set[OF sets] u by auto - also have "\ \ (AE x. u x \ 0)" by (auto split: split_max) + using AE_iff_null[OF sets] u by auto + also have "\ \ (AE x in M. u x \ 0)" by (auto split: split_max) finally show ?thesis . qed -lemma (in measure_space) positive_integral_const_If: - "(\\<^isup>+x. a \M) = (if 0 \ a then a * \ (space M) else 0)" +lemma positive_integral_const_If: + "(\\<^isup>+x. a \M) = (if 0 \ a then a * (emeasure M) (space M) else 0)" by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) -lemma (in measure_space) positive_integral_restricted: - assumes A: "A \ sets M" - 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 - - interpret R: measure_space ?R - by (rule restricted_measure_space) fact - let ?I = "\g x. g x * indicator A x :: ereal" - show ?thesis - unfolding positive_integral_def - unfolding simple_function_restricted[OF A] - unfolding AE_restricted[OF A] - proof (safe intro!: SUPR_eq) - fix g assume g: "simple_function M (?I g)" and le: "g \ max 0 \ f" - show "\j\{g. simple_function M g \ g \ max 0 \ ?I f}. - integral\<^isup>S (restricted_space A) g \ integral\<^isup>S M j" - proof (safe intro!: bexI[of _ "?I g"]) - show "integral\<^isup>S (restricted_space A) g \ integral\<^isup>S M (?I g)" - using g A by (simp add: simple_integral_restricted) - show "?I g \ max 0 \ ?I f" - using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) - qed fact - next - fix g assume g: "simple_function M g" and le: "g \ max 0 \ ?I f" - show "\i\{g. simple_function M (?I g) \ g \ max 0 \ f}. - integral\<^isup>S M g \ integral\<^isup>S (restricted_space A) i" - proof (safe intro!: bexI[of _ "?I g"]) - show "?I g \ max 0 \ f" - using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) - from le have "\x. g x \ ?I (?I g) x" - by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) - then show "integral\<^isup>S M g \ integral\<^isup>S (restricted_space A) (?I g)" - using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted) - show "simple_function M (?I (?I g))" using g A by auto - qed - qed -qed - -lemma (in measure_space) positive_integral_subalgebra: +lemma positive_integral_subalgebra: assumes f: "f \ borel_measurable N" "AE x in N. 0 \ f x" - and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" - and sa: "sigma_algebra N" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - interpret N: measure_space N using measure_space_subalgebra[OF sa N] . - from N.borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this + from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this note sf = simple_function_subalgebra[OF fs(1) N(1,2)] - from N.positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric] - have "integral\<^isup>P N f = (SUP i. \x\fs i ` space M. x * N.\ (fs i -` {x} \ space M))" + from positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric] + have "integral\<^isup>P N f = (SUP i. \x\fs i ` space M. x * emeasure N (fs i -` {x} \ space M))" unfolding fs(4) positive_integral_max_0 unfolding 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 fs(1)] unfolding `space N = space M` by auto + also have "\ = (SUP i. \x\fs i ` space M. x * (emeasure M) (fs i -` {x} \ space M))" + using N simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto also have "\ = integral\<^isup>P M f" using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric] unfolding fs(4) positive_integral_max_0 @@ -1673,7 +1422,7 @@ section "Lebesgue Integral" -definition integrable where +definition integrable :: "'a measure \ ('a \ real) \ bool" where "integrable M f \ f \ borel_measurable M \ (\\<^isup>+ x. ereal (f x) \M) \ \ \ (\\<^isup>+ x. ereal (- f x) \M) \ \" @@ -1682,55 +1431,44 @@ shows "f \ borel_measurable M" "(\\<^isup>+ x. ereal (f x) \M) \ \" "(\\<^isup>+ x. ereal (- f x) \M) \ \" using assms unfolding integrable_def by auto -definition lebesgue_integral_def: +definition lebesgue_integral :: "'a measure \ ('a \ real) \ real" ("integral\<^isup>L") where "integral\<^isup>L M f = real ((\\<^isup>+ x. ereal (f x) \M)) - real ((\\<^isup>+ x. ereal (- f x) \M))" syntax - "_lebesgue_integral" :: "pttrn \ real \ ('a, 'b) measure_space_scheme \ real" ("\ _. _ \_" [60,61] 110) + "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\ _. _ \_" [60,61] 110) translations - "\ x. f \M" == "CONST integral\<^isup>L M (%x. f)" + "\ x. f \M" == "CONST lebesgue_integral M (%x. f)" -lemma (in measure_space) integrableE: +lemma integrableE: assumes "integrable M f" obtains r q where "(\\<^isup>+x. ereal (f x)\M) = ereal r" "(\\<^isup>+x. ereal (-f x)\M) = ereal q" "f \ borel_measurable M" "integral\<^isup>L M f = r - q" using assms unfolding integrable_def lebesgue_integral_def - using positive_integral_positive[of "\x. ereal (f x)"] - using positive_integral_positive[of "\x. ereal (-f x)"] + using positive_integral_positive[of M "\x. ereal (f x)"] + using positive_integral_positive[of M "\x. ereal (-f x)"] by (cases rule: ereal2_cases[of "(\\<^isup>+x. ereal (-f x)\M)" "(\\<^isup>+x. ereal (f x)\M)"]) auto -lemma (in measure_space) integral_cong: +lemma integral_cong: 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 \ measure N A = \ A" "sets N = sets M" "space N = space M" - shows "integral\<^isup>L N f = integral\<^isup>L M f" - by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def) - -lemma (in measure_space) integrable_cong_measure: - assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" - shows "integrable N f \ integrable M f" - using assms - by (simp add: positive_integral_cong_measure[OF assms] integrable_def measurable_def) - -lemma (in measure_space) integral_cong_AE: - assumes cong: "AE x. f x = g x" +lemma integral_cong_AE: + assumes cong: "AE x in M. f x = g x" shows "integral\<^isup>L M f = integral\<^isup>L M g" proof - - have *: "AE x. ereal (f x) = ereal (g x)" - "AE x. ereal (- f x) = ereal (- g x)" using cong by auto + have *: "AE x in M. ereal (f x) = ereal (g x)" + "AE x in M. ereal (- f x) = ereal (- g x)" using cong by auto show ?thesis unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def .. qed -lemma (in measure_space) integrable_cong_AE: +lemma integrable_cong_AE: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" - assumes "AE x. f x = g x" + assumes "AE x in M. f x = g x" shows "integrable M f = integrable M g" proof - have "(\\<^isup>+ x. ereal (f x) \M) = (\\<^isup>+ x. ereal (g x) \M)" @@ -1740,11 +1478,23 @@ by (auto simp: integrable_def) qed -lemma (in measure_space) integrable_cong: +lemma integrable_cong: "(\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: +lemma positive_integral_eq_integral: + assumes f: "integrable M f" + assumes nonneg: "AE x in M. 0 \ f x" + shows "(\\<^isup>+ x. ereal (f x) \M) = integral\<^isup>L M f" +proof - + have "(\\<^isup>+ x. max 0 (ereal (- f x)) \M) = (\\<^isup>+ x. 0 \M)" + using nonneg by (intro positive_integral_cong_AE) (auto split: split_max) + with f positive_integral_positive show ?thesis + by (cases "\\<^isup>+ x. ereal (f x) \M") + (auto simp add: lebesgue_integral_def positive_integral_max_0 integrable_def) +qed + +lemma integral_eq_positive_integral: assumes f: "\x. 0 \ f x" shows "integral\<^isup>L M f = real (\\<^isup>+ x. ereal (f x) \M)" proof - @@ -1755,84 +1505,12 @@ unfolding lebesgue_integral_def by simp qed -lemma (in measure_space) integral_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - assumes f: "f \ borel_measurable M'" - shows "integral\<^isup>L M' f = (\x. f (T x) \M)" -proof - - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) - from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel] - have borel: "(\x. ereal (f x)) \ borel_measurable M'" "(\x. ereal (- f x)) \ borel_measurable M'" - and "(\x. f (T x)) \ borel_measurable M" - using f by (auto simp: comp_def) - then show ?thesis - using f unfolding lebesgue_integral_def integrable_def - by (auto simp: borel[THEN positive_integral_vimage[OF T]]) -qed - -lemma (in measure_space) integrable_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - assumes f: "integrable M' f" - shows "integrable M (\x. f (T x))" -proof - - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) - from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel] - have borel: "(\x. ereal (f x)) \ borel_measurable M'" "(\x. ereal (- f x)) \ borel_measurable M'" - and "(\x. f (T x)) \ borel_measurable M" - using f by (auto simp: comp_def) - then show ?thesis - using f unfolding lebesgue_integral_def integrable_def - by (auto simp: borel[THEN positive_integral_vimage[OF T]]) -qed - -lemma (in measure_space) integral_translated_density: - assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" - and g: "g \ borel_measurable M" - and N: "space N = space M" "sets N = sets M" - and density: "\A. A \ sets M \ measure N A = (\\<^isup>+ x. f x * indicator A x \M)" - (is "\A. _ \ _ = ?d A") - shows "integral\<^isup>L N g = (\ x. f x * g x \M)" (is ?integral) - and "integrable N g = integrable M (\x. f x * g x)" (is ?integrable) -proof - - from f have ms: "measure_space (M\measure := ?d\)" - by (intro measure_space_density[where u="\x. ereal (f x)"]) auto - - from ms density N have "(\\<^isup>+ x. g x \N) = (\\<^isup>+ x. max 0 (ereal (g x)) \M\measure := ?d\)" - unfolding positive_integral_max_0 - by (intro measure_space.positive_integral_cong_measure) auto - also have "\ = (\\<^isup>+ x. ereal (f x) * max 0 (ereal (g x)) \M)" - using f g by (intro positive_integral_translated_density) auto - also have "\ = (\\<^isup>+ x. max 0 (ereal (f x * g x)) \M)" - using f by (intro positive_integral_cong_AE) - (auto simp: ereal_max_0 zero_le_mult_iff split: split_max) - finally have pos: "(\\<^isup>+ x. g x \N) = (\\<^isup>+ x. f x * g x \M)" - by (simp add: positive_integral_max_0) - - from ms density N have "(\\<^isup>+ x. - (g x) \N) = (\\<^isup>+ x. max 0 (ereal (- g x)) \M\measure := ?d\)" - unfolding positive_integral_max_0 - by (intro measure_space.positive_integral_cong_measure) auto - also have "\ = (\\<^isup>+ x. ereal (f x) * max 0 (ereal (- g x)) \M)" - using f g by (intro positive_integral_translated_density) auto - also have "\ = (\\<^isup>+ x. max 0 (ereal (- f x * g x)) \M)" - using f by (intro positive_integral_cong_AE) - (auto simp: ereal_max_0 mult_le_0_iff split: split_max) - finally have neg: "(\\<^isup>+ x. - g x \N) = (\\<^isup>+ x. - (f x * g x) \M)" - by (simp add: positive_integral_max_0) - - have g_N: "g \ borel_measurable N" - using g N unfolding measurable_def by simp - - show ?integral ?integrable - unfolding lebesgue_integral_def integrable_def - using pos neg f g g_N by auto -qed - -lemma (in measure_space) integral_minus[intro, simp]: +lemma integral_minus[intro, simp]: 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_minus_iff[simp]: +lemma integral_minus_iff[simp]: "integrable M (\x. - f x) \ integrable M f" proof assume "integrable M (\x. - f x)" @@ -1841,7 +1519,7 @@ then show "integrable M f" by simp qed (rule integral_minus) -lemma (in measure_space) integral_of_positive_diff: +lemma integral_of_positive_diff: 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 M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" @@ -1851,7 +1529,7 @@ let ?u = "\x. max 0 (ereal (u x))" let ?v = "\x. max 0 (ereal (v x))" - from borel_measurable_diff[of u v] integrable + from borel_measurable_diff[of u M v] integrable have f_borel: "?f \ borel_measurable M" and mf_borel: "?mf \ borel_measurable M" and v_borel: "?v \ borel_measurable M" and @@ -1881,7 +1559,7 @@ using integrable f by (auto elim!: integrableE) qed -lemma (in measure_space) integral_linear: +lemma integral_linear: 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" (is ?EQ) @@ -1917,18 +1595,18 @@ by (auto elim!: integrableE simp: field_simps) qed -lemma (in measure_space) integral_add[simp, intro]: +lemma integral_add[simp, intro]: 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]: +lemma integral_zero[simp, intro]: 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]: +lemma integral_cmult[simp, intro]: 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) @@ -1942,37 +1620,80 @@ assume "a \ 0" hence "0 \ - a" by auto have *: "\t. - a * t + 0 = (-a) * t" by simp show ?thesis using integral_linear[OF assms integral_zero(1) `0 \ - a`] - integral_minus(1)[of "\t. - a * f t"] + integral_minus(1)[of M "\t. - a * f t"] unfolding * integral_zero by simp qed thus ?P ?I by auto qed -lemma (in measure_space) integral_multc: +lemma lebesgue_integral_cmult_nonneg: + assumes f: "f \ borel_measurable M" and "0 \ c" + shows "(\x. c * f x \M) = c * integral\<^isup>L M f" +proof - + { have "c * real (integral\<^isup>P M (\x. max 0 (ereal (f x)))) = + real (ereal c * integral\<^isup>P M (\x. max 0 (ereal (f x))))" + by simp + also have "\ = real (integral\<^isup>P M (\x. ereal c * max 0 (ereal (f x))))" + using f `0 \ c` by (subst positive_integral_cmult) auto + also have "\ = real (integral\<^isup>P M (\x. max 0 (ereal (c * f x))))" + using `0 \ c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def zero_le_mult_iff) + finally have "real (integral\<^isup>P M (\x. ereal (c * f x))) = c * real (integral\<^isup>P M (\x. ereal (f x)))" + by (simp add: positive_integral_max_0) } + moreover + { have "c * real (integral\<^isup>P M (\x. max 0 (ereal (- f x)))) = + real (ereal c * integral\<^isup>P M (\x. max 0 (ereal (- f x))))" + by simp + also have "\ = real (integral\<^isup>P M (\x. ereal c * max 0 (ereal (- f x))))" + using f `0 \ c` by (subst positive_integral_cmult) auto + also have "\ = real (integral\<^isup>P M (\x. max 0 (ereal (- c * f x))))" + using `0 \ c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def mult_le_0_iff) + finally have "real (integral\<^isup>P M (\x. ereal (- c * f x))) = c * real (integral\<^isup>P M (\x. ereal (- f x)))" + by (simp add: positive_integral_max_0) } + ultimately show ?thesis + by (simp add: lebesgue_integral_def field_simps) +qed + +lemma lebesgue_integral_uminus: + "(\x. - f x \M) = - integral\<^isup>L M f" + unfolding lebesgue_integral_def by simp + +lemma lebesgue_integral_cmult: + assumes f: "f \ borel_measurable M" + shows "(\x. c * f x \M) = c * integral\<^isup>L M f" +proof (cases rule: linorder_le_cases) + assume "0 \ c" with f show ?thesis by (rule lebesgue_integral_cmult_nonneg) +next + assume "c \ 0" + with lebesgue_integral_cmult_nonneg[OF f, of "-c"] + show ?thesis + by (simp add: lebesgue_integral_def) +qed + +lemma integral_multc: 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: +lemma integral_mono_AE: assumes fg: "integrable M f" "integrable M g" - and mono: "AE t. f t \ g t" + and mono: "AE t in M. f t \ g t" shows "integral\<^isup>L M f \ integral\<^isup>L M g" proof - - have "AE x. ereal (f x) \ ereal (g x)" + have "AE x in M. ereal (f x) \ ereal (g x)" using mono by auto - moreover have "AE x. ereal (- g x) \ ereal (- f x)" + moreover have "AE x in M. ereal (- g x) \ ereal (- f x)" using mono by auto ultimately show ?thesis using fg by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono simp: positive_integral_positive lebesgue_integral_def diff_minus) qed -lemma (in measure_space) integral_mono: +lemma integral_mono: assumes "integrable M f" "integrable M g" "\t. t \ space M \ f t \ g t" shows "integral\<^isup>L M f \ integral\<^isup>L M g" using assms by (auto intro: integral_mono_AE) -lemma (in measure_space) integral_diff[simp, intro]: +lemma integral_diff[simp, intro]: 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" @@ -1980,9 +1701,9 @@ 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\<^isup>L M (indicator A) = real (\ A)" (is ?int) +lemma integral_indicator[simp, intro]: + assumes "A \ sets M" and "(emeasure M) A \ \" + shows "integral\<^isup>L M (indicator A) = real ((emeasure M) A)" (is ?int) and "integrable M (indicator A)" (is ?able) proof - from `A \ sets M` have *: @@ -1994,10 +1715,10 @@ 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 \ \" +lemma integral_cmul_indicator: + assumes "A \ sets M" and "c \ 0 \ (emeasure M) A \ \" shows "integrable M (\x. c * indicator A x)" (is ?P) - and "(\x. c * indicator A x \M) = c * real (\ A)" (is ?I) + and "(\x. c * indicator A x \M) = c * real ((emeasure M) A)" (is ?I) proof - show ?P proof (cases "c = 0") @@ -2010,7 +1731,7 @@ qed simp qed -lemma (in measure_space) integral_setsum[simp, intro]: +lemma integral_setsum[simp, intro]: 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") @@ -2023,7 +1744,7 @@ thus "?int S" and "?I S" by auto qed -lemma (in measure_space) integrable_abs: +lemma integrable_abs: assumes "integrable M f" shows "integrable M (\ x. \f x\)" proof - @@ -2034,23 +1755,22 @@ by (simp add: positive_integral_add positive_integral_max_0 integrable_def) qed -lemma (in measure_space) integral_subalgebra: +lemma integral_subalgebra: assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" and sa: "sigma_algebra N" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" 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] . have "(\\<^isup>+ x. max 0 (ereal (f x)) \N) = (\\<^isup>+ x. max 0 (ereal (f x)) \M)" "(\\<^isup>+ x. max 0 (ereal (- f x)) \N) = (\\<^isup>+ x. max 0 (ereal (- f x)) \M)" - using borel by (auto intro!: positive_integral_subalgebra N sa) + using borel by (auto intro!: positive_integral_subalgebra N) moreover have "f \ borel_measurable M \ f \ borel_measurable N" using assms unfolding measurable_def by auto ultimately show ?P ?I by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0) qed -lemma (in measure_space) integrable_bound: +lemma integrable_bound: assumes "integrable M f" and f: "\x. x \ space M \ 0 \ f x" "\x. x \ space M \ \g x\ \ f x" @@ -2077,11 +1797,21 @@ unfolding integrable_def by auto qed -lemma (in measure_space) integrable_abs_iff: +lemma lebesgue_integral_nonneg: + assumes ae: "(AE x in M. 0 \ f x)" shows "0 \ integral\<^isup>L M f" +proof - + have "(\\<^isup>+x. max 0 (ereal (- f x)) \M) = (\\<^isup>+x. 0 \M)" + using ae by (intro positive_integral_cong_AE) (auto simp: max_def) + then show ?thesis + by (auto simp: lebesgue_integral_def positive_integral_max_0 + intro!: real_of_ereal_pos positive_integral_positive) +qed + +lemma integrable_abs_iff: "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: +lemma integrable_max: assumes int: "integrable M f" "integrable M g" shows "integrable M (\ x. max (f x) (g x))" proof (rule integrable_bound) @@ -2095,7 +1825,7 @@ by auto qed -lemma (in measure_space) integrable_min: +lemma integrable_min: assumes int: "integrable M f" "integrable M g" shows "integrable M (\ x. min (f x) (g x))" proof (rule integrable_bound) @@ -2109,18 +1839,18 @@ by auto qed -lemma (in measure_space) integral_triangle_inequality: +lemma integral_triangle_inequality: assumes "integrable M f" shows "\integral\<^isup>L M f\ \ (\x. \f x\ \M)" proof - 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] + using assms integral_minus(2)[of M f, symmetric] by (auto intro!: integral_mono integrable_abs simp del: integral_minus) finally show ?thesis . qed -lemma (in measure_space) integral_positive: +lemma integral_positive: assumes "integrable M f" "\x. x \ space M \ 0 \ f x" shows "0 \ integral\<^isup>L M f" proof - @@ -2130,7 +1860,7 @@ finally show ?thesis . qed -lemma (in measure_space) integral_monotone_convergence_pos: +lemma integral_monotone_convergence_pos: 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" @@ -2157,7 +1887,7 @@ using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def) have integral_eq: "\n. (\\<^isup>+ x. ereal (f n x) \M) = ereal (integral\<^isup>L M (f n))" - using i positive_integral_positive by (auto simp: ereal_real lebesgue_integral_def integrable_def) + using i positive_integral_positive[of M] by (auto simp: ereal_real lebesgue_integral_def integrable_def) have pos_integral: "\n. 0 \ integral\<^isup>L M (f n)" using pos i by (auto simp: integral_positive) @@ -2177,7 +1907,7 @@ unfolding integrable_def lebesgue_integral_def by auto qed -lemma (in measure_space) integral_monotone_convergence: +lemma integral_monotone_convergence: assumes f: "\i. integrable M (f i)" and "mono f" and lim: "\x. (\i. f i x) ----> u x" and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" @@ -2201,9 +1931,9 @@ by (auto simp: integral_diff) qed -lemma (in measure_space) integral_0_iff: +lemma integral_0_iff: assumes "integrable M f" - shows "(\x. \f x\ \M) = 0 \ \ {x\space M. f x \ 0} = 0" + shows "(\x. \f x\ \M) = 0 \ (emeasure M) {x\space M. f x \ 0} = 0" proof - have *: "(\\<^isup>+x. ereal (- \f x\) \M) = 0" using assms by (auto simp: positive_integral_0_iff_AE integrable_def) @@ -2212,57 +1942,57 @@ "(\\<^isup>+ x. ereal \f x\ \M) \ \" unfolding integrable_def by auto from positive_integral_0_iff[OF this(1)] this(2) show ?thesis unfolding lebesgue_integral_def * - using positive_integral_positive[of "\x. ereal \f x\"] + using positive_integral_positive[of M "\x. ereal \f x\"] by (auto simp add: real_of_ereal_eq_0) qed -lemma (in measure_space) positive_integral_PInf: +lemma positive_integral_PInf: assumes f: "f \ borel_measurable M" and not_Inf: "integral\<^isup>P M f \ \" - shows "\ (f -` {\} \ space M) = 0" + shows "(emeasure M) (f -` {\} \ space M) = 0" proof - - have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" + have "\ * (emeasure M) (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets) also have "\ \ integral\<^isup>P M (\x. max 0 (f x))" by (auto intro!: positive_integral_mono simp: indicator_def max_def) - finally have "\ * \ (f -` {\} \ space M) \ integral\<^isup>P M f" + finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^isup>P M f" by (simp add: positive_integral_max_0) - moreover have "0 \ \ (f -` {\} \ space M)" - using f by (simp add: measurable_sets) + moreover have "0 \ (emeasure M) (f -` {\} \ space M)" + by (rule emeasure_nonneg) ultimately show ?thesis using assms by (auto split: split_if_asm) qed -lemma (in measure_space) positive_integral_PInf_AE: - assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x. f x \ \" +lemma positive_integral_PInf_AE: + assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x in M. f x \ \" proof (rule AE_I) - show "\ (f -` {\} \ space M) = 0" + show "(emeasure M) (f -` {\} \ space M) = 0" by (rule positive_integral_PInf[OF assms]) show "f -` {\} \ space M \ sets M" using assms by (auto intro: borel_measurable_vimage) qed auto -lemma (in measure_space) simple_integral_PInf: +lemma simple_integral_PInf: assumes "simple_function M f" "\x. 0 \ f x" and "integral\<^isup>S M f \ \" - shows "\ (f -` {\} \ space M) = 0" + shows "(emeasure M) (f -` {\} \ space M) = 0" proof (rule positive_integral_PInf) show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) show "integral\<^isup>P M f \ \" using assms by (simp add: positive_integral_eq_simple_integral) qed -lemma (in measure_space) integral_real: - "AE x. \f x\ \ \ \ (\x. real (f x) \M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\x. - f x))" +lemma integral_real: + "AE x in M. \f x\ \ \ \ (\x. real (f x) \M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\x. - f x))" using assms unfolding lebesgue_integral_def by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real) lemma (in finite_measure) lebesgue_integral_const[simp]: shows "integrable M (\x. a)" - and "(\x. a \M) = a * \' (space M)" + and "(\x. a \M) = a * (measure M) (space M)" proof - { fix a :: real assume "0 \ a" - then have "(\\<^isup>+ x. ereal a \M) = ereal a * \ (space M)" + then have "(\\<^isup>+ x. ereal a \M) = ereal a * (emeasure M) (space M)" by (subst positive_integral_const) auto moreover from `0 \ a` have "(\\<^isup>+ x. ereal (-a) \M) = 0" @@ -2277,8 +2007,8 @@ then have "0 \ -a" by auto from *[OF this] show ?thesis by simp qed - show "(\x. a \M) = a * \' (space M)" - by (simp add: \'_def lebesgue_integral_def positive_integral_const_If) + show "(\x. a \M) = a * measure M (space M)" + by (simp add: lebesgue_integral_def positive_integral_const_If emeasure_eq_measure) qed lemma indicator_less[simp]: @@ -2287,8 +2017,8 @@ lemma (in finite_measure) integral_less_AE: assumes int: "integrable M X" "integrable M Y" - assumes A: "\ A \ 0" "A \ sets M" "AE x. x \ A \ X x \ Y x" - assumes gt: "AE x. X x \ Y x" + assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" + assumes gt: "AE x in M. X x \ Y x" shows "integral\<^isup>L M X < integral\<^isup>L M Y" proof - have "integral\<^isup>L M X \ integral\<^isup>L M Y" @@ -2301,26 +2031,27 @@ using gt by (intro integral_cong_AE) auto also have "\ = 0" using eq int by simp - finally have "\ {x \ space M. Y x - X x \ 0} = 0" + finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" using int by (simp add: integral_0_iff) moreover have "(\\<^isup>+x. indicator A x \M) \ (\\<^isup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" using A by (intro positive_integral_mono_AE) auto - then have "\ A \ \ {x \ space M. Y x - X x \ 0}" + then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" using int A by (simp add: integrable_def) - moreover note `\ A \ 0` positive_measure[OF `A \ sets M`] - ultimately show False by auto + ultimately have "emeasure M A = 0" + using emeasure_nonneg[of M A] by simp + with `(emeasure M) A \ 0` show False by auto qed ultimately show ?thesis by auto qed lemma (in finite_measure) integral_less_AE_space: assumes int: "integrable M X" "integrable M Y" - assumes gt: "AE x. X x < Y x" "\ (space M) \ 0" + assumes gt: "AE x in M. X x < Y x" "(emeasure M) (space M) \ 0" shows "integral\<^isup>L M X < integral\<^isup>L M Y" using gt by (intro integral_less_AE[OF int, where A="space M"]) auto -lemma (in measure_space) integral_dominated_convergence: +lemma integral_dominated_convergence: assumes u: "\i. integrable M (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" and w: "integrable M w" and u': "\x. x \ space M \ (\i. u i x) ----> u' x" @@ -2336,7 +2067,7 @@ from u[unfolded integrable_def] have u'_borel: "u' \ borel_measurable M" - using u' by (blast intro: borel_measurable_LIMSEQ[of u]) + using u' by (blast intro: borel_measurable_LIMSEQ[of M u]) { fix x assume x: "x \ space M" then have "0 \ \u 0 x\" by auto @@ -2385,7 +2116,7 @@ using diff_less_2w[of _ n] unfolding positive_integral_max_0 by (intro positive_integral_mono) auto then have "?f n = 0" - using positive_integral_positive[of ?f'] eq_0 by auto } + using positive_integral_positive[of M ?f'] eq_0 by auto } then show ?thesis by (simp add: Limsup_const) next assume neq_0: "(\\<^isup>+ x. max 0 (ereal (2 * w x)) \M) \ 0" (is "?wx \ 0") @@ -2411,10 +2142,10 @@ also have "\ = (\\<^isup>+ x. ereal (2 * w x) \M) - limsup (\n. \\<^isup>+ x. ereal \u n x - u' x\ \M)" unfolding PI_diff positive_integral_max_0 - using positive_integral_positive[of "\x. ereal (2 * w x)"] + using positive_integral_positive[of M "\x. ereal (2 * w x)"] by (subst liminf_ereal_cminus) auto finally show ?thesis - using neq_0 I2w_fin positive_integral_positive[of "\x. ereal (2 * w x)"] pos + using neq_0 I2w_fin positive_integral_positive[of M "\x. ereal (2 * w x)"] pos unfolding positive_integral_max_0 by (cases rule: ereal2_cases[of "\\<^isup>+ x. ereal (2 * w x) \M" "limsup (\n. \\<^isup>+ x. ereal \u n x - u' x\ \M)"]) auto @@ -2430,8 +2161,8 @@ ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0" using `limsup ?f = 0` by auto have "\n. (\\<^isup>+ x. ereal \u n x - u' x\ \M) = ereal (\x. \u n x - u' x\ \M)" - using diff positive_integral_positive - by (subst integral_eq_positive_integral) (auto simp: ereal_real integrable_def) + using diff positive_integral_positive[of M] + by (subst integral_eq_positive_integral[of _ M]) (auto simp: ereal_real integrable_def) then show ?lim_diff using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq] by (simp add: lim_ereal) @@ -2456,7 +2187,7 @@ qed qed -lemma (in measure_space) integral_sums: +lemma integral_sums: 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\ \M))" @@ -2513,18 +2244,18 @@ section "Lebesgue integration on countable spaces" -lemma (in measure_space) integral_on_countable: +lemma integral_on_countable: assumes f: "f \ borel_measurable M" and bij: "bij_betw enum S (f ` space M)" 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))\)" + and fin: "\x. x \ 0 \ (emeasure M) (f -` {x} \ space M) \ \" + and abs_summable: "summable (\r. \enum r * real ((emeasure M) (f -` {enum r} \ space M))\)" shows "integrable M f" - and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral\<^isup>L M f" (is ?sums) + and "(\r. enum r * real ((emeasure M) (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\<^isup>L M (?F r)" + have enum_eq: "\r. enum r * real ((emeasure M) (?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" @@ -2551,9 +2282,9 @@ { fix r 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))" + also have "\ = \enum r\ * real ((emeasure M) (?A r))" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - finally have "(\x. \?F r x\ \M) = \enum r * real (\ (?A r))\" + finally have "(\x. \?F r x\ \M) = \enum r * real ((emeasure M) (?A r))\" using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) } note int_abs_F = this @@ -2570,67 +2301,469 @@ show "integrable M f" unfolding int_f by simp qed -section "Lebesgue integration on finite space" +section {* Distributions *} + +lemma simple_function_distr[simp]: + "simple_function (distr M M' T) f \ simple_function M' (\x. f x)" + unfolding simple_function_def by simp + +lemma positive_integral_distr: + assumes T: "T \ measurable M M'" + and f: "f \ borel_measurable M'" + shows "integral\<^isup>P (distr M M' T) f = (\\<^isup>+ x. f (T x) \M)" +proof - + from borel_measurable_implies_simple_function_sequence'[OF f] + guess f' . note f' = this + then have f_distr: "\i. simple_function (distr M M' T) (f' i)" + by simp + let ?f = "\i x. f' i (T x)" + have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def) + have sup: "\x. (SUP i. ?f i x) = max 0 (f (T x))" + using f'(4) . + have sf: "\i. simple_function M (\x. f' i (T x))" + using simple_function_comp[OF T(1) f'(1)] . + show "integral\<^isup>P (distr M M' T) f = (\\<^isup>+ x. f (T x) \M)" + using + positive_integral_monotone_convergence_simple[OF f'(2,5) f_distr] + positive_integral_monotone_convergence_simple[OF inc f'(5) sf] + by (simp add: positive_integral_max_0 simple_integral_distr[OF T f'(1)] f') +qed + +lemma integral_distr: + assumes T: "T \ measurable M M'" + assumes f: "f \ borel_measurable M'" + shows "integral\<^isup>L (distr M M' T) f = (\x. f (T x) \M)" +proof - + from measurable_comp[OF T, of f borel] + have borel: "(\x. ereal (f x)) \ borel_measurable M'" "(\x. ereal (- f x)) \ borel_measurable M'" + and "(\x. f (T x)) \ borel_measurable M" + using f by (auto simp: comp_def) + then show ?thesis + using f unfolding lebesgue_integral_def integrable_def + by (auto simp: borel[THEN positive_integral_distr[OF T]]) +qed -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 M f" - and "(\x. f x \M) = - (\ r \ f`space M. r * real (\ (f -` {r} \ space M)))" (is "?integral") +lemma integrable_distr: + assumes T: "T \ measurable M M'" and f: "integrable (distr M M' T) f" + shows "integrable M (\x. f (T x))" +proof - + from measurable_comp[OF T, of f borel] + have borel: "(\x. ereal (f x)) \ borel_measurable M'" "(\x. ereal (- f x)) \ borel_measurable M'" + and "(\x. f (T x)) \ borel_measurable M" + using f by (auto simp: comp_def) + then show ?thesis + using f unfolding lebesgue_integral_def integrable_def + using borel[THEN positive_integral_distr[OF T]] + by (auto simp: borel[THEN positive_integral_distr[OF T]]) +qed + +section {* Lebesgue integration on @{const count_space} *} + +lemma simple_function_count_space[simp]: + "simple_function (count_space A) f \ finite (f ` A)" + unfolding simple_function_def by simp + +lemma positive_integral_count_space: + assumes A: "finite {a\A. 0 < f a}" + shows "integral\<^isup>P (count_space A) f = (\a|a\A \ 0 < f a. f a)" proof - - let ?A = "\r. f -` {r} \ space M" - let ?S = "\x. \r\f`space M. r * indicator (?A r) x" + have *: "(\\<^isup>+x. max 0 (f x) \count_space A) = + (\\<^isup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" + by (auto intro!: positive_integral_cong + simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less) + also have "\ = (\a|a\A \ 0 < f a. \\<^isup>+ x. f a * indicator {a} x \count_space A)" + by (subst positive_integral_setsum) + (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) + also have "\ = (\a|a\A \ 0 < f a. f a)" + by (auto intro!: setsum_cong simp: positive_integral_cmult_indicator one_ereal_def[symmetric]) + finally show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma integrable_count_space: + "finite X \ integrable (count_space X) f" + by (auto simp: positive_integral_count_space integrable_def) + +lemma positive_integral_count_space_finite: + "finite A \ (\\<^isup>+x. f x \count_space A) = (\a\A. max 0 (f a))" + by (subst positive_integral_max_0[symmetric]) + (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le) + +lemma lebesgue_integral_count_space_finite: + "finite A \ (\x. f x \count_space A) = (\a\A. f a)" + apply (auto intro!: setsum_mono_zero_left + simp: positive_integral_count_space_finite lebesgue_integral_def) + apply (subst (1 2) setsum_real_of_ereal[symmetric]) + apply (auto simp: max_def setsum_subtractf[symmetric] intro!: setsum_cong) + done + +section {* Measure spaces with an associated density *} + +definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where + "density M f = measure_of (space M) (sets M) (\A. \\<^isup>+ x. f x * indicator A x \M)" - { fix x assume "x \ space M" - have "f x = (\r\f`space M. if x \ ?A r then r else 0)" - using finite `x \ space M` by (simp add: setsum_cases) - also have "\ = ?S x" - by (auto intro!: setsum_cong) - finally have "f x = ?S x" . } - note f_eq = this +lemma + shows sets_density[simp]: "sets (density M f) = sets M" + and space_density[simp]: "space (density M f) = space M" + by (auto simp: density_def) + +lemma + shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" + and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" + and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" + unfolding measurable_def simple_function_def by simp_all + +lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ + (AE x in M. f x = f' x) \ density M f = density M f'" + unfolding density_def by (auto intro!: measure_of_eq positive_integral_cong_AE space_closed) + +lemma density_max_0: "density M f = density M (\x. max 0 (f x))" +proof - + have "\x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" + by (auto simp: indicator_def) + then show ?thesis + unfolding density_def by (simp add: positive_integral_max_0) +qed + +lemma density_ereal_max_0: "density M (\x. ereal (f x)) = density M (\x. ereal (max 0 (f x)))" + by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) - 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) +lemma emeasure_density: + assumes f: "f \ borel_measurable M" and A: "A \ sets M" + shows "emeasure (density M f) A = (\\<^isup>+ x. f x * indicator A x \M)" + (is "_ = ?\ A") + unfolding density_def +proof (rule emeasure_measure_of_sigma) + show "sigma_algebra (space M) (sets M)" .. + show "positive (sets M) ?\" + using f by (auto simp: positive_def intro!: positive_integral_positive) + have \_eq: "?\ = (\A. \\<^isup>+ x. max 0 (f x) * indicator A x \M)" (is "?\ = ?\'") + apply (subst positive_integral_max_0[symmetric]) + apply (intro ext positive_integral_cong_AE AE_I2) + apply (auto simp: indicator_def) + done + show "countably_additive (sets M) ?\" + unfolding \_eq + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume "range A \ sets M" + then have *: "\i. (\x. max 0 (f x) * indicator (A i) x) \ borel_measurable M" + using f by (auto intro!: borel_measurable_ereal_times) + assume disj: "disjoint_family A" + have "(\n. ?\' (A n)) = (\\<^isup>+ x. (\n. max 0 (f x) * indicator (A n) x) \M)" + using f * by (simp add: positive_integral_suminf) + also have "\ = (\\<^isup>+ x. max 0 (f x) * (\n. indicator (A n) x) \M)" using f + by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE) + also have "\ = (\\<^isup>+ x. max 0 (f x) * indicator (\n. A n) x \M)" + unfolding suminf_indicator[OF disj] .. + finally show "(\n. ?\' (A n)) = ?\' (\x. A x)" by simp + qed +qed fact - show "integrable M f" ?integral using fin f f_eq_S - by (simp_all add: integral_cmul_indicator borel_measurable_vimage) +lemma null_sets_density_iff: + assumes f: "f \ borel_measurable M" + shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x \ 0)" +proof - + { assume "A \ sets M" + have eq: "(\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. max 0 (f x) * indicator A x \M)" + apply (subst positive_integral_max_0[symmetric]) + apply (intro positive_integral_cong) + apply (auto simp: indicator_def) + done + have "(\\<^isup>+x. f x * indicator A x \M) = 0 \ + emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" + unfolding eq + using f `A \ sets M` + by (intro positive_integral_0_iff) auto + also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" + using f `A \ sets M` + by (intro AE_iff_measurable[OF _ refl, symmetric]) + (auto intro!: sets_Collect borel_measurable_ereal_eq) + also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" + by (auto simp add: indicator_def max_def split: split_if_asm) + finally have "(\\<^isup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } + with f show ?thesis + by (simp add: null_sets_def emeasure_density cong: conj_cong) +qed + +lemma AE_density: + assumes f: "f \ borel_measurable M" + shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" +proof + assume "AE x in density M f. P x" + with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x \ 0" + by (auto simp: eventually_ae_filter null_sets_density_iff) + then have "AE x in M. x \ N \ P x" by auto + with ae show "AE x in M. 0 < f x \ P x" + by (rule eventually_elim2) auto +next + fix N assume ae: "AE x in M. 0 < f x \ P x" + then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" + by (auto simp: eventually_ae_filter) + then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. \ 0 < f x}" + "N \ {x\space M. \ 0 < f x} \ sets M" and ae2: "AE x in M. x \ N" + using f by (auto simp: subset_eq intro!: sets_Collect_neg AE_not_in) + show "AE x in density M f. P x" + using ae2 + unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] + by (intro exI[of _ "N \ {x\space M. \ 0 < f x}"] conjI *) + (auto elim: eventually_elim2) qed -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 positive_integral_density: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g': "g' \ borel_measurable M" + shows "integral\<^isup>P (density M f) g' = (\\<^isup>+ x. f x * g' x \M)" +proof - + def g \ "\x. max 0 (g' x)" + then have g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + using g' by auto + from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this + note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)] + note G'(2)[simp] + { fix P have "AE x in M. P x \ AE x in M. P x" + using positive_integral_null_set[of _ _ f] + by (auto simp: eventually_ae_filter ) } + note ac = this + with G(4) f g have G_M': "AE x in density M f. (SUP i. G i x) = g x" + by (auto simp add: AE_density[OF f(1)] max_def) + { fix i + let ?I = "\y x. indicator (G i -` {y} \ space M) x" + { fix x assume *: "x \ space M" "0 \ f x" "0 \ g x" + then have [simp]: "G i ` space M \ {y. G i x = y \ x \ space M} = {G i x}" by auto + from * G' G have "(\y\G i`space M. y * (f x * ?I y x)) = f x * (\y\G i`space M. (y * ?I y x))" + by (subst setsum_ereal_right_distrib) (auto simp: ac_simps) + also have "\ = f x * G i x" + by (simp add: indicator_def if_distrib setsum_cases) + finally have "(\y\G i`space M. y * (f x * ?I y x)) = f x * G i x" . } + note to_singleton = this + have "integral\<^isup>P (density M f) (G i) = integral\<^isup>S (density M f) (G i)" + using G by (intro positive_integral_eq_simple_integral) simp_all + also have "\ = (\y\G i`space M. y * (\\<^isup>+x. f x * ?I y x \M))" + using f G(1) + by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_density + simp: simple_function_def simple_integral_def) + also have "\ = (\y\G i`space M. (\\<^isup>+x. y * (f x * ?I y x) \M))" + using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric]) + also have "\ = (\\<^isup>+x. (\y\G i`space M. y * (f x * ?I y x)) \M)" + using f G' G by (auto intro!: positive_integral_setsum[symmetric]) + finally have "integral\<^isup>P (density M f) (G i) = (\\<^isup>+x. f x * G i x \M)" + using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) } + note [simp] = this + have "integral\<^isup>P (density M f) g = (SUP i. integral\<^isup>P (density M f) (G i))" using G'(1) G_M'(1) G + using positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`, of "density M f"] + by (simp cong: positive_integral_cong_AE) + also have "\ = (SUP i. (\\<^isup>+x. f x * G i x \M))" by simp + also have "\ = (\\<^isup>+x. (SUP i. f x * G i x) \M)" + using f G' G(2)[THEN incseq_SucD] G + by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) + (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff) + also have "\ = (\\<^isup>+x. f x * g x \M)" using f G' G g + by (intro positive_integral_cong_AE) + (auto simp add: SUPR_ereal_cmult split: split_max) + also have "\ = (\\<^isup>+x. f x * g' x \M)" + using f(2) + by (subst (2) positive_integral_max_0[symmetric]) + (auto simp: g_def max_def ereal_zero_le_0_iff intro!: positive_integral_cong_AE) + finally show "integral\<^isup>P (density M f) g' = (\\<^isup>+x. f x * g' x \M)" + unfolding g_def positive_integral_max_0 . +qed -lemma (in finite_measure_space) positive_integral_finite_eq_setsum: - assumes pos: "\x. x \ space M \ 0 \ f x" - shows "integral\<^isup>P M f = (\x \ space M. f x * \ {x})" +lemma integral_density: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and g: "g \ borel_measurable M" + shows "integral\<^isup>L (density M f) g = (\ x. f x * g x \M)" + and "integrable (density M f) g \ integrable M (\x. f x * g x)" + unfolding lebesgue_integral_def integrable_def using f g + by (auto simp: positive_integral_density) + +lemma emeasure_restricted: + assumes S: "S \ sets M" and X: "X \ sets M" + shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" proof - - 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] pos - by (simp add: positive_integral_setsum positive_integral_cmult_indicator) + have "emeasure (density M (indicator S)) X = (\\<^isup>+x. indicator S x * indicator X x \M)" + using S X by (simp add: emeasure_density) + also have "\ = (\\<^isup>+x. indicator (S \ X) x \M)" + by (auto intro!: positive_integral_cong simp: indicator_def) + also have "\ = emeasure M (S \ X)" + using S X by (simp add: Int) + finally show ?thesis . +qed + +lemma measure_restricted: + "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" + by (simp add: emeasure_restricted measure_def) + +lemma (in finite_measure) finite_measure_restricted: + "S \ sets M \ finite_measure (density M (indicator S))" + by default (simp add: emeasure_restricted) + +lemma emeasure_density_const: + "A \ sets M \ 0 \ c \ emeasure (density M (\_. c)) A = c * emeasure M A" + by (auto simp: positive_integral_cmult_indicator emeasure_density) + +lemma measure_density_const: + "A \ sets M \ 0 < c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" + by (auto simp: emeasure_density_const measure_def) + +lemma density_density_eq: + "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. 0 \ f x \ + density (density M f) g = density M (\x. f x * g x)" + by (auto intro!: measure_eqI simp: emeasure_density positive_integral_density ac_simps) + +lemma distr_density_distr: + assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" + and inv: "\x\space M. T' (T x) = x" + assumes f: "f \ borel_measurable M'" + shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") +proof (rule measure_eqI) + fix A assume A: "A \ sets ?R" + { fix x assume "x \ space M" + with sets_into_space[OF A] + have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ereal)" + using T inv by (auto simp: indicator_def measurable_space) } + with A T T' f show "emeasure ?R A = emeasure ?L A" + by (simp add: measurable_comp emeasure_density emeasure_distr + positive_integral_distr measurable_sets cong: positive_integral_cong) +qed simp + +lemma density_density_divide: + fixes f g :: "'a \ real" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + assumes ac: "AE x in M. f x = 0 \ g x = 0" + shows "density (density M f) (\x. g x / f x) = density M g" +proof - + have "density M g = density M (\x. f x * (g x / f x))" + using f g ac by (auto intro!: density_cong measurable_If) + then show ?thesis + using f g by (subst density_density_eq) auto qed -lemma (in finite_measure_space) integral_finite_singleton: - shows "integrable M f" - and "integral\<^isup>L M f = (\x \ space M. f x * real (\ {x}))" (is ?I) +section {* Point measure *} + +definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where + "point_measure A f = density (count_space A) f" + +lemma + shows space_point_measure: "space (point_measure A f) = A" + and sets_point_measure: "sets (point_measure A f) = Pow A" + by (auto simp: point_measure_def) + +lemma measurable_point_measure_eq1[simp]: + "g \ measurable (point_measure A f) M \ g \ A \ space M" + unfolding point_measure_def by simp + +lemma measurable_point_measure_eq2_finite[simp]: + "finite A \ + g \ measurable M (point_measure A f) \ + (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" + unfolding point_measure_def by simp + +lemma simple_function_point_measure[simp]: + "simple_function (point_measure A f) g \ finite (g ` A)" + by (simp add: point_measure_def) + +lemma emeasure_point_measure: + assumes A: "finite {a\X. 0 < f a}" "X \ A" + shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" proof - - have *: - "(\\<^isup>+ x. max 0 (ereal (f x)) \M) = (\x \ space M. max 0 (ereal (f x)) * \ {x})" - "(\\<^isup>+ x. max 0 (ereal (- f x)) \M) = (\x \ space M. max 0 (ereal (- f x)) * \ {x})" - by (simp_all add: positive_integral_finite_eq_setsum) - then show "integrable M f" using finite_space finite_measure - by (simp add: setsum_Pinfty integrable_def positive_integral_max_0 - split: split_max) - show ?I using finite_measure * - apply (simp add: positive_integral_max_0 lebesgue_integral_def) - apply (subst (1 2) setsum_real_of_ereal[symmetric]) - apply (simp_all split: split_max add: setsum_subtractf[symmetric]) - apply (intro setsum_cong[OF refl]) - apply (simp split: split_max) - done + have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" + using `X \ A` by auto + with A show ?thesis + by (simp add: emeasure_density positive_integral_count_space ereal_zero_le_0_iff + point_measure_def indicator_def) qed +lemma emeasure_point_measure_finite: + "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a|a\X. f a)" + by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) + +lemma null_sets_point_measure_iff: + "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x \ 0)" + by (auto simp: AE_count_space null_sets_density_iff point_measure_def) + +lemma AE_point_measure: + "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" + unfolding point_measure_def + by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) + +lemma positive_integral_point_measure: + "finite {a\A. 0 < f a \ 0 < g a} \ + integral\<^isup>P (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" + unfolding point_measure_def + apply (subst density_max_0) + apply (subst positive_integral_density) + apply (simp_all add: AE_count_space positive_integral_density) + apply (subst positive_integral_count_space ) + apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff) + apply (rule finite_subset) + prefer 2 + apply assumption + apply auto + done + +lemma positive_integral_point_measure_finite: + "finite A \ (\a. a \ A \ 0 \ f a) \ (\a. a \ A \ 0 \ g a) \ + integral\<^isup>P (point_measure A f) g = (\a\A. f a * g a)" + by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) + +lemma lebesgue_integral_point_measure_finite: + "finite A \ (\a. a \ A \ 0 \ f a) \ integral\<^isup>L (point_measure A f) g = (\a\A. f a * g a)" + by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) + +lemma integrable_point_measure_finite: + "finite A \ integrable (point_measure A (\x. ereal (f x))) g" + unfolding point_measure_def + apply (subst density_ereal_max_0) + apply (subst integral_density) + apply (auto simp: AE_count_space integrable_count_space) + done + +section {* Uniform measure *} + +definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" + +lemma + shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" + and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" + by (auto simp: uniform_measure_def) + +lemma emeasure_uniform_measure[simp]: + assumes A: "A \ sets M" and B: "B \ sets M" + shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" +proof - + from A B have "emeasure (uniform_measure M A) B = (\\<^isup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" + by (auto simp add: uniform_measure_def emeasure_density split: split_indicator + intro!: positive_integral_cong) + also have "\ = emeasure M (A \ B) / emeasure M A" + using A B + by (subst positive_integral_cmult_indicator) (simp_all add: Int emeasure_nonneg) + finally show ?thesis . +qed + +lemma emeasure_neq_0_sets: "emeasure M A \ 0 \ A \ sets M" + using emeasure_notin_sets[of A M] by blast + +lemma measure_uniform_measure[simp]: + assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" + shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" + using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A + by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) + +section {* Uniform count measure *} + +definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" + +lemma + shows space_uniform_count_measure: "space (uniform_count_measure A) = A" + and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" + unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) + +lemma emeasure_uniform_count_measure: + "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) + +lemma measure_uniform_count_measure: + "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) + end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -9,6 +9,15 @@ imports Finite_Product_Measure begin +lemma borel_measurable_sets: + assumes "f \ measurable borel M" "A \ sets M" + shows "f -` A \ sets borel" + using measurable_sets[OF assms] by simp + +lemma measurable_identity[intro,simp]: + "(\x. x) \ measurable M M" + unfolding measurable_def by auto + subsection {* Standard Cubes *} definition cube :: "nat \ 'a::ordered_euclidean_space set" where @@ -52,17 +61,13 @@ subsection {* Lebesgue measure *} -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. ereal (integral (cube n) (indicator A)) \" +definition lebesgue :: "'a::ordered_euclidean_space measure" where + "lebesgue = measure_of UNIV {A. \n. (indicator A :: 'a \ real) integrable_on cube n} + (\A. SUP n. ereal (integral (cube n) (indicator A)))" lemma space_lebesgue[simp]: "space lebesgue = UNIV" unfolding lebesgue_def by simp -lemma lebesgueD: "A \ sets lebesgue \ (indicator A :: _ \ real) integrable_on cube n" - unfolding lebesgue_def by simp - lemma lebesgueI: "(\n. (indicator A :: _ \ real) integrable_on cube n) \ A \ sets lebesgue" unfolding lebesgue_def by simp @@ -86,23 +91,23 @@ "A \ B = {} \ (indicator A x::_::monoid_add) + indicator B x = indicator (A \ B) x" unfolding indicator_def by auto -interpretation lebesgue: sigma_algebra lebesgue -proof (intro sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI lebesgueI) - fix A n assume A: "A \ sets lebesgue" - have "indicator (space lebesgue - A) = (\x. 1 - indicator A x :: real)" +lemma sigma_algebra_lebesgue: + defines "leb \ {A. \n. (indicator A :: 'a::ordered_euclidean_space \ real) integrable_on cube n}" + shows "sigma_algebra UNIV leb" +proof (safe intro!: sigma_algebra_iff2[THEN iffD2]) + fix A assume A: "A \ leb" + moreover have "indicator (UNIV - A) = (\x. 1 - indicator A x :: real)" by (auto simp: fun_eq_iff indicator_def) - then show "(indicator (space lebesgue - A) :: _ \ real) integrable_on cube n" - using A by (auto intro!: integrable_sub dest: lebesgueD simp: cube_def) + ultimately show "UNIV - A \ leb" + using A by (auto intro!: integrable_sub simp: cube_def leb_def) next - fix n show "(indicator {} :: _\real) integrable_on cube n" - by (auto simp: cube_def indicator_def [abs_def]) + fix n show "{} \ leb" + by (auto simp: cube_def indicator_def[abs_def] leb_def) next - fix A :: "nat \ 'a set" and n ::nat assume "range A \ sets lebesgue" - then have A: "\i. (indicator (A i) :: _ \ real) integrable_on cube n" - by (auto dest: lebesgueD) - show "(indicator (\i. A i) :: _ \ real) integrable_on cube n" (is "?g integrable_on _") - proof (intro dominated_convergence[where g="?g"] ballI) - fix k show "(indicator (\i real) integrable_on cube n" + fix A :: "nat \ _" assume A: "range A \ leb" + have "\n. (indicator (\i. A i) :: _ \ real) integrable_on cube n" (is "\n. ?g integrable_on _") + proof (intro dominated_convergence[where g="?g"] ballI allI) + fix k n show "(indicator (\i real) integrable_on cube n" proof (induct k) case (Suc k) have *: "(\ i i A k" @@ -111,36 +116,45 @@ indicator (\ ix. max (?f x) (?g x)) = _") by (auto simp: fun_eq_iff * indicator_def) show ?case - using absolutely_integrable_max[of ?f "cube n" ?g] A Suc by (simp add: *) + using absolutely_integrable_max[of ?f "cube n" ?g] A Suc + by (simp add: * leb_def subset_eq) qed auto qed (auto intro: LIMSEQ_indicator_UN simp: cube_def) + then show "(\i. A i) \ leb" by (auto simp: leb_def) qed simp -interpretation lebesgue: measure_space lebesgue -proof +lemma sets_lebesgue: "sets lebesgue = {A. \n. (indicator A :: _ \ real) integrable_on cube n}" + unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] .. + +lemma lebesgueD: "A \ sets lebesgue \ (indicator A :: _ \ real) integrable_on cube n" + unfolding sets_lebesgue by simp + +lemma emeasure_lebesgue: + assumes "A \ sets lebesgue" + shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))" + (is "_ = ?\ A") +proof (rule emeasure_measure_of[OF lebesgue_def]) have *: "indicator {} = (\x. 0 :: real)" by (simp add: fun_eq_iff) - show "positive lebesgue (measure lebesgue)" - proof (unfold positive_def, safe) - show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def) - fix A assume "A \ sets lebesgue" - then show "0 \ measure lebesgue A" - unfolding lebesgue_def - by (auto intro!: SUP_upper2 integral_nonneg) + show "positive (sets lebesgue) ?\" + proof (unfold positive_def, intro conjI ballI) + show "?\ {} = 0" by (simp add: integral_0 *) + fix A :: "'a set" assume "A \ sets lebesgue" then show "0 \ ?\ A" + by (auto intro!: SUP_upper2 Integration.integral_nonneg simp: sets_lebesgue) qed next - show "countably_additive lebesgue (measure lebesgue)" + show "countably_additive (sets lebesgue) ?\" proof (intro countably_additive_def[THEN iffD2] allI impI) - fix A :: "nat \ 'b set" assume rA: "range A \ sets lebesgue" "disjoint_family A" + fix A :: "nat \ 'a 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" by (auto dest: lebesgueD) let ?m = "\n i. integral (cube n) (indicator (A i) :: _\real)" let ?M = "\n I. integral (cube n) (indicator (\i\I. A i) :: _\real)" - have nn[simp, intro]: "\i n. 0 \ ?m n i" by (auto intro!: integral_nonneg) + have nn[simp, intro]: "\i n. 0 \ ?m n i" by (auto intro!: Integration.integral_nonneg) 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 "(\n. measure lebesgue (A n)) = measure lebesgue (\i. A i)" - proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI) + by (auto simp: sets_lebesgue) + show "(\n. ?\ (A n)) = ?\ (\i. A i)" + proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) fix i n show "ereal (?m n i) \ ereal (?m (Suc n) i)" using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI) next @@ -172,14 +186,15 @@ indicator (\im. \x = 0.. ?M n UNIV" by (simp add: atLeast0LessThan) qed qed qed -qed +next +qed (auto, fact) lemma has_integral_interval_cube: fixes a b :: "'a::ordered_euclidean_space" @@ -202,9 +217,10 @@ fixes s::"'a::ordered_euclidean_space set" assumes "s \ sets borel" shows "s \ sets lebesgue" proof - - let ?S = "range (\(a, b). {a .. (b :: 'a\ordered_euclidean_space)})" - have *:"?S \ sets lebesgue" - proof (safe intro!: lebesgueI) + have "s \ sigma_sets (space lebesgue) (range (\(a, b). {a .. (b :: 'a\ordered_euclidean_space)}))" + using assms by (simp add: borel_eq_atLeastAtMost) + also have "\ \ sets lebesgue" + proof (safe intro!: sigma_sets_subset lebesgueI) fix n :: nat and a b :: 'a let ?N = "\\ i. max (- real n) (a $$ i)" let ?P = "\\ i. min (real n) (b $$ i)" @@ -212,11 +228,7 @@ unfolding integrable_on_def using has_integral_interval_cube[of a b] by auto qed - have "s \ sigma_sets UNIV ?S" using assms - unfolding borel_eq_atLeastAtMost by (simp add: sigma_def) - thus ?thesis - using lebesgue.sigma_subset[of "\ space = UNIV, sets = ?S\", simplified, OF *] - by (auto simp: sigma_def) + finally show ?thesis . qed lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set" @@ -224,19 +236,21 @@ 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 "lebesgue.\ S = 0" + fixes S :: "'a::ordered_euclidean_space set" + assumes "negligible S" shows "emeasure lebesgue S = 0" proof - have "\n. integral (cube n) (indicator S :: 'a\real) = 0" 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) + then show ?thesis + using assms by (simp add: emeasure_lebesgue lebesgueI_negligible) qed lemma lmeasure_iff_LIMSEQ: - assumes "A \ sets lebesgue" "0 \ m" - shows "lebesgue.\ A = ereal m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" -proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ) + assumes A: "A \ sets lebesgue" and "0 \ m" + shows "emeasure lebesgue A = ereal m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" +proof (subst emeasure_lebesgue[OF A], 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) qed @@ -261,7 +275,7 @@ lemma lmeasure_finite_has_integral: fixes s :: "'a::ordered_euclidean_space set" - assumes "s \ sets lebesgue" "lebesgue.\ s = ereal m" "0 \ m" + assumes "s \ sets lebesgue" "emeasure lebesgue s = ereal m" "0 \ m" shows "(indicator s has_integral m) UNIV" proof - let ?I = "indicator :: 'a set \ 'a \ real" @@ -275,7 +289,7 @@ (auto dest!: lebesgueD) } moreover { fix n have "0 \ integral (cube n) (?I s)" - using assms by (auto dest!: lebesgueD intro!: integral_nonneg) } + using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) } ultimately show "bounded {integral UNIV (?I (s \ cube k)) |k. True}" unfolding bounded_def @@ -303,14 +317,13 @@ unfolding m by (intro integrable_integral **) qed -lemma lmeasure_finite_integrable: assumes s: "s \ sets lebesgue" and "lebesgue.\ s \ \" +lemma lmeasure_finite_integrable: assumes s: "s \ sets lebesgue" and "emeasure lebesgue s \ \" shows "(indicator s :: _ \ real) integrable_on UNIV" -proof (cases "lebesgue.\ s") +proof (cases "emeasure lebesgue s") case (real m) - with lmeasure_finite_has_integral[OF `s \ sets lebesgue` this] - lebesgue.positive_measure[OF s] + with lmeasure_finite_has_integral[OF `s \ sets lebesgue` this] emeasure_nonneg[of lebesgue s] show ?thesis unfolding integrable_on_def by auto -qed (insert assms lebesgue.positive_measure[OF s], auto) +qed (insert assms emeasure_nonneg[of lebesgue s], auto) lemma has_integral_lebesgue: assumes "((indicator s :: _\real) has_integral m) UNIV" shows "s \ sets lebesgue" @@ -324,7 +337,7 @@ qed lemma has_integral_lmeasure: assumes "((indicator s :: _\real) has_integral m) UNIV" - shows "lebesgue.\ s = ereal m" + shows "emeasure lebesgue s = ereal 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] . @@ -349,55 +362,56 @@ qed lemma has_integral_iff_lmeasure: - "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = ereal m)" + "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ emeasure lebesgue A = ereal 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 \ lebesgue.\ A = ereal m" + show "A \ sets lebesgue \ 0 \ m \ emeasure lebesgue A = ereal m" by (auto intro: has_integral_nonneg) next - assume "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = ereal m" + assume "A \ sets lebesgue \ 0 \ m \ emeasure lebesgue A = ereal 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 "lebesgue.\ s = ereal (integral UNIV (indicator s))" + shows "emeasure lebesgue s = ereal (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 "lebesgue.\ s = ereal (integral UNIV (indicator s))" by simp + show "emeasure lebesgue s = ereal (integral UNIV (indicator s))" by simp qed lemma lebesgue_simple_function_indicator: fixes f::"'a::ordered_euclidean_space \ ereal" assumes f:"simple_function lebesgue f" shows "f = (\x. (\y \ f ` UNIV. y * indicator (f -` {y}) x))" - by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto + by (rule, subst simple_function_indicator_representation[OF f]) auto lemma integral_eq_lmeasure: - "(indicator s::_\real) integrable_on UNIV \ integral UNIV (indicator s) = real (lebesgue.\ s)" + "(indicator s::_\real) integrable_on UNIV \ integral UNIV (indicator s) = real (emeasure lebesgue s)" by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg) -lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "lebesgue.\ s \ \" +lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "emeasure lebesgue s \ \" using lmeasure_eq_integral[OF assms] by auto lemma negligible_iff_lebesgue_null_sets: - "negligible A \ A \ lebesgue.null_sets" + "negligible A \ A \ null_sets lebesgue" proof assume "negligible A" from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0] - show "A \ lebesgue.null_sets" by auto + show "A \ null_sets lebesgue" by auto next - assume A: "A \ lebesgue.null_sets" - then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A] by auto + assume A: "A \ null_sets lebesgue" + then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A] + by (auto simp: null_sets_def) show "negligible A" unfolding negligible_def proof (intro allI) fix a b :: 'a have integrable: "(indicator A :: _\real) integrable_on {a..b}" by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *) then have "integral {a..b} (indicator A) \ (integral UNIV (indicator A) :: real)" - using * by (auto intro!: integral_subset_le has_integral_integrable) + using * by (auto intro!: integral_subset_le) moreover have "(0::real) \ integral {a..b} (indicator A)" using integrable by (auto intro!: integral_nonneg) ultimately have "integral {a..b} (indicator A) = (0::real)" @@ -412,8 +426,8 @@ shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) -lemma lmeasure_UNIV[intro]: "lebesgue.\ (UNIV::'a::ordered_euclidean_space set) = \" -proof (simp add: lebesgue_def, intro SUP_PInfty bexI) +lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \" +proof (simp add: emeasure_lebesgue, intro SUP_PInfty bexI) fix n :: nat have "indicator UNIV = (\x::'a. 1 :: real)" by auto moreover @@ -434,7 +448,7 @@ lemma fixes a b ::"'a::ordered_euclidean_space" - shows lmeasure_atLeastAtMost[simp]: "lebesgue.\ {a..b} = ereal (content {a..b})" + shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})" proof - have "(indicator (UNIV \ {a..b})::_\real) integrable_on UNIV" unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def]) @@ -454,7 +468,7 @@ qed lemma lmeasure_singleton[simp]: - fixes a :: "'a::ordered_euclidean_space" shows "lebesgue.\ {a} = 0" + fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0" using lmeasure_atLeastAtMost[of a a] by simp declare content_real[simp] @@ -462,82 +476,68 @@ lemma fixes a b :: real shows lmeasure_real_greaterThanAtMost[simp]: - "lebesgue.\ {a <.. b} = ereal (if a \ b then b - a else 0)" + "emeasure lebesgue {a <.. b} = ereal (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lebesgue.\ {a <.. b} = lebesgue.\ {a .. b} - lebesgue.\ {a}" - by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lebesgue.\]) + then have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b} - emeasure lebesgue {a}" + by (subst emeasure_Diff[symmetric]) + (auto intro!: arg_cong[where f="emeasure lebesgue"]) then show ?thesis by auto qed auto lemma fixes a b :: real shows lmeasure_real_atLeastLessThan[simp]: - "lebesgue.\ {a ..< b} = ereal (if a \ b then b - a else 0)" + "emeasure lebesgue {a ..< b} = ereal (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lebesgue.\ {a ..< b} = lebesgue.\ {a .. b} - lebesgue.\ {b}" - by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lebesgue.\]) + then have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b} - emeasure lebesgue {b}" + by (subst emeasure_Diff[symmetric]) + (auto intro!: arg_cong[where f="emeasure lebesgue"]) then show ?thesis by auto qed auto lemma fixes a b :: real shows lmeasure_real_greaterThanLessThan[simp]: - "lebesgue.\ {a <..< b} = ereal (if a \ b then b - a else 0)" + "emeasure lebesgue {a <..< b} = ereal (if a \ b then b - a else 0)" proof cases assume "a < b" - then have "lebesgue.\ {a <..< b} = lebesgue.\ {a <.. b} - lebesgue.\ {b}" - by (subst lebesgue.measure_Diff[symmetric]) - (auto intro!: arg_cong[where f=lebesgue.\]) + then have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a <.. b} - emeasure lebesgue {b}" + by (subst emeasure_Diff[symmetric]) + (auto intro!: arg_cong[where f="emeasure lebesgue"]) then show ?thesis by auto qed auto subsection {* Lebesgue-Borel measure *} -definition "lborel = lebesgue \ sets := sets borel \" +definition "lborel = measure_of UNIV (sets borel) (emeasure lebesgue)" 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 [abs_def] lborel_def) + and measurable_lborel1[simp]: "measurable lborel = measurable borel" + and measurable_lborel2[simp]: "measurable A lborel = measurable A borel" + using sigma_sets_eq[of borel] + by (auto simp add: lborel_def measurable_def[abs_def]) -interpretation lborel: measure_space "lborel :: ('a::ordered_euclidean_space) measure_space" - where "space lborel = UNIV" - and "sets lborel = sets borel" - and "measure lborel = lebesgue.\" - and "measurable lborel = measurable borel" -proof (rule lebesgue.measure_space_subalgebra) - have "sigma_algebra (lborel::'a measure_space) \ sigma_algebra (borel::'a algebra)" - unfolding sigma_algebra_iff2 lborel_def by simp - then show "sigma_algebra (lborel::'a measure_space)" by simp default -qed auto +lemma emeasure_lborel[simp]: "A \ sets borel \ emeasure lborel A = emeasure lebesgue A" + by (rule emeasure_measure_of[OF lborel_def]) + (auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure) 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 +proof (default, intro conjI exI[of _ "\n. cube n"]) + show "range cube \ sets lborel" by (auto intro: borel_closed) + { fix x :: 'a have "\n. x\cube n" using mem_big_cube by auto } + then show "(\i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto + show "\i. emeasure lborel (cube i) \ \" by (simp add: cube_def) +qed interpretation lebesgue: sigma_finite_measure lebesgue proof - 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. lebesgue.\ (A i) \ \)" - by auto + from lborel.sigma_finite guess A :: "nat \ 'a set" .. + then show "\A::nat \ 'a set. range A \ sets lebesgue \ (\i. A i) = space lebesgue \ (\i. emeasure lebesgue (A i) \ \)" + by (intro exI[of _ A]) (auto simp: subset_eq) qed subsection {* Lebesgue integrable implies Gauge integrable *} @@ -556,11 +556,11 @@ fixes f::"'a::ordered_euclidean_space \ ereal" assumes f:"simple_function lebesgue f" and f':"range f \ {0..<\}" - and om:"\x. x \ range f \ lebesgue.\ (f -` {x} \ UNIV) = \ \ x = 0" + and om:"\x. x \ range f \ emeasure lebesgue (f -` {x} \ UNIV) = \ \ x = 0" shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" unfolding simple_integral_def space_lebesgue proof (subst lebesgue_simple_function_indicator) - let ?M = "\x. lebesgue.\ (f -` {x} \ UNIV)" + let ?M = "\x. emeasure lebesgue (f -` {x} \ UNIV)" let ?F = "\x. indicator (f -` {x})" { fix x y assume "y \ range f" from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)" @@ -571,7 +571,7 @@ have "x * ?M x = real x * real (?M x)" proof cases assume "x \ 0" with om[OF x] have "?M x \ \" by auto - with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis + with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis by (cases rule: ereal2_cases[of x "?M x"]) auto qed simp } ultimately @@ -580,11 +580,11 @@ by simp also have \ proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral - real_of_ereal_pos lebesgue.positive_measure ballI) - show *: "finite (range f)" "\y. f -` {y} \ sets lebesgue" "\y. f -` {y} \ UNIV \ sets lebesgue" - using lebesgue.simple_functionD[OF f] by auto + real_of_ereal_pos emeasure_nonneg ballI) + show *: "finite (range f)" "\y. f -` {y} \ sets lebesgue" + using simple_functionD[OF f] by auto fix y assume "real y \ 0" "y \ range f" - with * om[OF this(2)] show "lebesgue.\ (f -` {y}) = ereal (real (?M y))" + with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))" by (auto simp: ereal_real) qed finally show "((\x. real (\y\range f. y * ?F y x)) has_integral real (\x\range f. x * ?M x)) UNIV" . @@ -601,28 +601,28 @@ shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" proof - let ?f = "\x. if x \ f -` {\} then 0 else f x" - note f(1)[THEN lebesgue.simple_functionD(2)] + note f(1)[THEN simple_functionD(2)] then have [simp, intro]: "\X. f -` X \ sets lebesgue" by auto have f': "simple_function lebesgue ?f" - using f by (intro lebesgue.simple_function_If_set) auto + using f by (intro simple_function_If_set) auto have rng: "range ?f \ {0..<\}" using f by auto have "AE x in lebesgue. f x = ?f x" - using lebesgue.simple_integral_PInf[OF f i] - by (intro lebesgue.AE_I[where N="f -` {\} \ space lebesgue"]) auto + using simple_integral_PInf[OF f i] + by (intro AE_I[where N="f -` {\} \ space lebesgue"]) auto from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f" - by (rule lebesgue.simple_integral_cong_AE) + by (rule simple_integral_cong_AE) have real_eq: "\x. real (f x) = real (?f x)" by auto show ?thesis unfolding eq real_eq proof (rule simple_function_has_integral[OF f' rng]) - fix x assume x: "x \ range ?f" and inf: "lebesgue.\ (?f -` {x} \ UNIV) = \" - have "x * lebesgue.\ (?f -` {x} \ UNIV) = (\\<^isup>S y. x * indicator (?f -` {x}) y \lebesgue)" - using f'[THEN lebesgue.simple_functionD(2)] - by (simp add: lebesgue.simple_integral_cmult_indicator) + fix x assume x: "x \ range ?f" and inf: "emeasure lebesgue (?f -` {x} \ UNIV) = \" + have "x * emeasure lebesgue (?f -` {x} \ UNIV) = (\\<^isup>S y. x * indicator (?f -` {x}) y \lebesgue)" + using f'[THEN simple_functionD(2)] + by (simp add: simple_integral_cmult_indicator) also have "\ \ integral\<^isup>S lebesgue f" - using f'[THEN lebesgue.simple_functionD(2)] f - by (intro lebesgue.simple_integral_mono lebesgue.simple_function_mult lebesgue.simple_function_indicator) + using f'[THEN simple_functionD(2)] f + by (intro simple_integral_mono simple_function_mult simple_function_indicator) (auto split: split_indicator) finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm) qed @@ -633,16 +633,16 @@ assumes f: "f \ borel_measurable lebesgue" "range f \ {0..<\}" "integral\<^isup>P lebesgue f \ \" shows "((\x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV" proof - - from lebesgue.borel_measurable_implies_simple_function_sequence'[OF f(1)] + from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . note u = this have SUP_eq: "\x. (SUP i. u i x) = f x" using u(4) f(2)[THEN subsetD] by (auto split: split_max) let ?u = "\i x. real (u i x)" - note u_eq = lebesgue.positive_integral_eq_simple_integral[OF u(1,5), symmetric] + note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric] { fix i note u_eq also have "integral\<^isup>P lebesgue (u i) \ (\\<^isup>+x. max 0 (f x) \lebesgue)" - by (intro lebesgue.positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric]) + by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric]) finally have "integral\<^isup>S lebesgue (u i) \ \" unfolding positive_integral_max_0 using f by auto } note u_fin = this @@ -684,10 +684,10 @@ also have "\ = real (integral\<^isup>S lebesgue (u k))" using u_int[THEN integral_unique] by (simp add: u') also have "\ = real (integral\<^isup>P lebesgue (u k))" - using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp + using positive_integral_eq_simple_integral[OF u(1,5)] by simp also have "\ \ real (integral\<^isup>P lebesgue f)" using f - by (auto intro!: real_of_ereal_positive_mono lebesgue.positive_integral_positive - lebesgue.positive_integral_mono SUP_upper simp: SUP_eq[symmetric]) + by (auto intro!: real_of_ereal_positive_mono positive_integral_positive + positive_integral_mono SUP_upper simp: SUP_eq[symmetric]) finally show "\integral UNIV (u' k)\ \ real (integral\<^isup>P lebesgue f)" . qed qed @@ -695,21 +695,21 @@ have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')" proof (rule tendsto_unique[OF trivial_limit_sequentially]) have "(\i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))" - unfolding u_eq by (intro LIMSEQ_ereal_SUPR lebesgue.incseq_positive_integral u) - also note lebesgue.positive_integral_monotone_convergence_SUP - [OF u(2) lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric] + unfolding u_eq by (intro LIMSEQ_ereal_SUPR incseq_positive_integral u) + also note positive_integral_monotone_convergence_SUP + [OF u(2) borel_measurable_simple_function[OF u(1)] u(5), symmetric] finally show "(\k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f" unfolding SUP_eq . { fix k have "0 \ integral\<^isup>S lebesgue (u k)" - using u by (auto intro!: lebesgue.simple_integral_positive) + using u by (auto intro!: simple_integral_positive) then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))" using u_fin by (auto simp: ereal_real) } note * = this show "(\k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')" using convergent using u_int[THEN integral_unique, symmetric] - by (subst *) (simp add: lim_ereal u') + by (subst *) (simp add: u') qed then show ?thesis using convergent by (simp add: f' integrable_integral) qed @@ -721,8 +721,8 @@ proof - let ?n = "\x. real (ereal (max 0 (- f x)))" and ?p = "\x. real (ereal (max 0 (f x)))" have *: "f = (\x. ?p x - ?n x)" by (auto simp del: ereal_max) - { fix f have "(\\<^isup>+ x. ereal (f x) \lebesgue) = (\\<^isup>+ x. ereal (max 0 (f x)) \lebesgue)" - by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) } + { fix f :: "'a \ real" have "(\\<^isup>+ x. ereal (f x) \lebesgue) = (\\<^isup>+ x. ereal (max 0 (f x)) \lebesgue)" + by (intro positive_integral_cong_pos) (auto split: split_max) } note eq = this show ?thesis unfolding lebesgue_integral_def @@ -732,7 +732,7 @@ apply (safe intro!: positive_integral_has_integral) using integrableD[OF f] by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0 split: split_max - intro!: lebesgue.measurable_If lebesgue.borel_measurable_ereal) + intro!: measurable_If) qed lemma lebesgue_positive_integral_eq_borel: @@ -740,7 +740,7 @@ shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f" proof - from f have "integral\<^isup>P lebesgue (\x. max 0 (f x)) = integral\<^isup>P lborel (\x. max 0 (f x))" - by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default + by (auto intro!: positive_integral_subalgebra[symmetric]) then show ?thesis unfolding positive_integral_max_0 . qed @@ -749,9 +749,8 @@ 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 lborel" by default have "sets lborel \ sets lebesgue" by auto - from lebesgue.integral_subalgebra[of f lborel, OF _ this _ _ *] assms + from integral_subalgebra[of f lborel, OF _ this _ _] assms show ?P ?I by auto qed @@ -783,152 +782,109 @@ "p2e (e2p x) = (x::'a::ordered_euclidean_space)" by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def) -interpretation lborel_product: product_sigma_finite "\x. lborel::real measure_space" +interpretation lborel_product: product_sigma_finite "\x. lborel::real measure" by default -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) {..x. lborel::real measure" "{..x\A. \y. P x y) \ (\f. \x\A. P x (f x))" + by metis 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. {..\<^isub>M i\I. lborel) = sigma_sets (\\<^isub>E i\I. UNIV) { \\<^isub>E i\I. {..< x i :: real} | x. True}" (is "_ = ?G") +proof (subst sigma_prod_algebra_sigma_eq[where S="\_ i::nat. {..M I (\i. lborel))) {Pi\<^isub>E I F |F. \i\I. F i \ range lessThan} = ?G" + by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff) +qed (auto simp: borel_eq_lessThan incseq_def reals_Archimedean2 image_iff intro: real_natceiling_ge) lemma measurable_e2p: - "e2p \ measurable (borel::'a::ordered_euclidean_space 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 (borel::'a::ordered_euclidean_space measure) (\\<^isub>M i\{.. real) set" assume "A \ {\\<^isub>E i\{..\<^isub>E i\{..\ i. x i) :: 'a}" + using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def + euclidean_eq[where 'a='a] eucl_less[where 'a='a]) + then show "e2p -` A \ space (borel::'a measure) \ sets borel" by simp +qed (auto simp: e2p_def) lemma measurable_p2e: - "p2e \ measurable (\\<^isub>M i\{.. 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 = {.. space ?P = (\\<^isub>E i\{.. x}" + assume "i < DIM('a)" + then have "?A = (\\<^isub>E j\{.. sets ?P" + by auto +qed + +lemma Int_stable_atLeastAtMost: + fixes x::"'a::ordered_euclidean_space" + shows "Int_stable (range (\(a, b::'a). {a..b}))" + by (auto simp: inter_interval Int_stable_def) -lemma Int_stable_cuboids: - fixes x::"'a::ordered_euclidean_space" - shows "Int_stable \space = UNIV, sets = range (\(a, b::'a). {a..b})\" - by (auto simp: inter_interval Int_stable_def) +lemma lborel_eqI: + fixes M :: "'a::ordered_euclidean_space measure" + assumes emeasure_eq: "\a b. emeasure M {a .. b} = content {a .. b}" + assumes sets_eq: "sets M = sets borel" + shows "lborel = M" +proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost]) + let ?P = "\\<^isub>M i\{.. Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E" + by (simp_all add: borel_eq_atLeastAtMost sets_eq) + + show "range cube \ ?E" unfolding cube_def [abs_def] by auto + show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI) + { fix x :: 'a have "\n. x \ cube n" using mem_big_cube[of x] by fastforce } + then show "(\i. cube i :: 'a set) = UNIV" by auto + + { fix i show "emeasure lborel (cube i) \ \" unfolding cube_def by auto } + { fix X assume "X \ ?E" then show "emeasure lborel X = emeasure M X" + by (auto simp: emeasure_eq) } +qed lemma lborel_eq_lborel_space: - fixes A :: "('a::ordered_euclidean_space) set" - assumes "A \ sets borel" - shows "lborel.\ A = lborel_space.\ DIM('a) (p2e -` A \ (space (lborel_space.P DIM('a))))" - (is "_ = measure ?P (?T A)") -proof (rule measure_unique_Int_stable_vimage) - show "measure_space ?P" by default - show "measure_space lborel" by default - - let ?E = "\ space = UNIV :: 'a set, sets = range (\(a,b). {a..b}) \" - show "Int_stable ?E" using Int_stable_cuboids . - show "range cube \ sets ?E" unfolding cube_def [abs_def] by auto - show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI) - { fix x have "\n. x \ cube n" using mem_big_cube[of x] by fastforce } - then show "(\i. cube i) = space ?E" by auto - { fix i show "lborel.\ (cube i) \ \" unfolding cube_def by auto } - show "A \ sets (sigma ?E)" "sets (sigma ?E) = sets lborel" "space ?E = space lborel" - using assms by (simp_all add: borel_eq_atLeastAtMost) - - show "p2e \ measurable ?P (lborel :: 'a measure_space)" - using measurable_p2e unfolding measurable_def by simp - { fix X assume "X \ sets ?E" - then obtain a b where X[simp]: "X = {a .. b}" by auto - have *: "?T X = (\\<^isub>E i\{.. {a $$ x .. b $$ x})" - by (auto simp: content_closed_interval eucl_le[where 'a='a] - intro!: setprod_ereal[symmetric]) - also have "\ = measure ?P (?T X)" - unfolding * by (subst lborel_space.measure_times) auto - finally show ?thesis . - qed simp } + "(lborel :: 'a measure) = distr (\\<^isub>M i\{.. space ?P = (\\<^isub>E i\{.. measure_preserving ?P ?E") -proof - show "p2e \ measurable ?P ?E" - using measurable_p2e by (simp add: measurable_def) - fix A :: "'a set" assume "A \ sets lborel" - then show "lborel_space.\ DIM('a) (p2e -` A \ (space (lborel_space.P DIM('a)))) = lborel.\ A" - by (intro lborel_eq_lborel_space[symmetric]) simp -qed - -lemma lebesgue_eq_lborel_space_in_borel: - fixes A :: "('a::ordered_euclidean_space) set" - assumes A: "A \ sets borel" - shows "lebesgue.\ A = lborel_space.\ DIM('a) (p2e -` A \ (space (lborel_space.P DIM('a))))" - using lborel_eq_lborel_space[OF A] by simp - lemma borel_fubini_positiv_integral: fixes f :: "'a::ordered_euclidean_space \ ereal" assumes f: "f \ borel_measurable borel" - shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(lborel_space.P DIM('a))" -proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _]) - show "f \ borel_measurable lborel" - using f by (simp_all add: measurable_def) -qed default + shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(\\<^isub>M i\{.. real" shows "integrable lborel f \ - integrable (lborel_space.P DIM('a)) (\x. f (p2e x))" + integrable (\\<^isub>M i\{..x. f (p2e x))" (is "_ \ integrable ?B ?f") proof assume "integrable lborel f" @@ -941,9 +897,9 @@ by (simp add: comp_def borel_fubini_positiv_integral integrable_def) next assume "integrable ?B ?f" - moreover then - have "?f \ e2p \ borel_measurable (borel::'a algebra)" - by (auto intro!: measurable_e2p measurable_comp) + moreover + then have "?f \ e2p \ borel_measurable (borel::'a measure)" + by (auto intro!: measurable_e2p) then have "f \ borel_measurable borel" by (simp cong: measurable_cong) ultimately show "integrable lborel f" @@ -953,100 +909,35 @@ lemma borel_fubini: fixes f :: "'a::ordered_euclidean_space \ real" assumes f: "f \ borel_measurable borel" - shows "integral\<^isup>L lborel f = \x. f (p2e x) \(lborel_space.P DIM('a))" + shows "integral\<^isup>L lborel f = \x. f (p2e x) \((\\<^isub>M i\{..space = UNIV, sets = range (\(a,b). {a::'a::ordered_euclidean_space .. b}) \" -proof (simp add: Int_stable_def image_iff, intro allI) - fix a1 b1 a2 b2 :: 'a - have "\ia b. {a1$$i..b1$$i} \ {a2$$i..b2$$i} = {a..b}" by auto - then have "\a b. \i {a2$$i..b2$$i} = {a i..b i}" - unfolding choice_iff' . - then guess a b by safe - then have "{a1..b1} \ {a2..b2} = {(\\ i. a i) .. (\\ i. b i)}" - by (simp add: set_eq_iff eucl_le[where 'a='a] all_conj_distrib[symmetric]) blast - then show "\a' b'. {a1..b1} \ {a2..b2} = {a'..b'}" by blast -qed - -lemma (in sigma_algebra) borel_measurable_sets: - assumes "f \ measurable borel M" "A \ sets M" - shows "f -` A \ sets borel" - using measurable_sets[OF assms] by simp - -lemma (in sigma_algebra) measurable_identity[intro,simp]: - "(\x. x) \ measurable M M" - unfolding measurable_def by auto +lemma borel_measurable_indicator': + "A \ sets borel \ f \ borel_measurable M \ (\x. indicator A (f x)) \ borel_measurable M" + using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def) lemma lebesgue_real_affine: - fixes X :: "real set" - assumes "X \ sets borel" and "c \ 0" - shows "measure lebesgue X = ereal \c\ * measure lebesgue ((\x. t + c * x) -` X)" - (is "_ = ?\ X") -proof - - let ?E = "\space = UNIV, sets = range (\(a,b). {a::real .. b})\ :: real algebra" - let ?M = "\\. \space = space ?E, sets = sets (sigma ?E), measure = \\ :: real measure_space" - have *: "?M (measure lebesgue) = lborel" - unfolding borel_eq_atLeastAtMost[symmetric] - by (simp add: lborel_def lebesgue_def) - have **: "?M ?\ = lborel \ measure := ?\ \" - unfolding borel_eq_atLeastAtMost[symmetric] - by (simp add: lborel_def lebesgue_def) - show ?thesis - proof (rule measure_unique_Int_stable[where X=X, OF Int_stable_atLeastAtMost], unfold * **) - show "X \ sets (sigma ?E)" - unfolding borel_eq_atLeastAtMost[symmetric] by fact - have "\x. \xa. x \ cube xa" apply(rule_tac x=x in mem_big_cube) by fastforce - then show "(\i. cube i) = space ?E" by auto - show "incseq cube" by (intro incseq_SucI cube_subset_Suc) - show "range cube \ sets ?E" - unfolding cube_def [abs_def] by auto - show "\i. measure lebesgue (cube i) \ \" - by (simp add: cube_def) - show "measure_space lborel" by default - then interpret sigma_algebra "lborel\measure := ?\\" - by (auto simp add: measure_space_def) - show "measure_space (lborel\measure := ?\\)" - proof - show "positive (lborel\measure := ?\\) (measure (lborel\measure := ?\\))" - by (auto simp: positive_def intro!: ereal_0_le_mult borel.borel_measurable_sets) - show "countably_additive (lborel\measure := ?\\) (measure (lborel\measure := ?\\))" - proof (simp add: countably_additive_def, safe) - fix A :: "nat \ real set" assume A: "range A \ sets borel" "disjoint_family A" - then have Ai: "\i. A i \ sets borel" by auto - have "(\n. measure lebesgue ((\x. t + c * x) -` A n)) = measure lebesgue (\n. (\x. t + c * x) -` A n)" - proof (intro lborel.measure_countably_additive) - { fix n have "(\x. t + c * x) -` A n \ space borel \ sets borel" - using A borel.measurable_ident unfolding id_def - by (intro measurable_sets[where A=borel] borel.borel_measurable_add[OF _ borel.borel_measurable_times]) auto } - then show "range (\i. (\x. t + c * x) -` A i) \ sets borel" by auto - from `disjoint_family A` - show "disjoint_family (\i. (\x. t + c * x) -` A i)" - by (rule disjoint_family_on_bisimulation) auto - qed - with Ai show "(\n. ?\ (A n)) = ?\ (UNION UNIV A)" - by (subst suminf_cmult_ereal) - (auto simp: vimage_UN borel.borel_measurable_sets) - qed - qed - fix X assume "X \ sets ?E" - then obtain a b where [simp]: "X = {a .. b}" by auto - show "measure lebesgue X = ?\ X" - proof cases - assume "0 < c" - then have "(\x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}" - by (auto simp: field_simps) - with `0 < c` show ?thesis - by (cases "a \ b") (auto simp: field_simps) - next - assume "\ 0 < c" with `c \ 0` have "c < 0" by auto - then have *: "(\x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}" - by (auto simp: field_simps) - with `c < 0` show ?thesis - by (cases "a \ b") (auto simp: field_simps) - qed + fixes c :: real assumes "c \ 0" + shows "lborel = density (distr lborel borel (\x. t + c * x)) (\_. \c\)" (is "_ = ?D") +proof (rule lborel_eqI) + fix a b show "emeasure ?D {a..b} = content {a .. b}" + proof cases + assume "0 < c" + then have "(\x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}" + by (auto simp: field_simps) + with `0 < c` show ?thesis + by (cases "a \ b") + (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult + borel_measurable_indicator' emeasure_distr) + next + assume "\ 0 < c" with `c \ 0` have "c < 0" by auto + then have *: "(\x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}" + by (auto simp: field_simps) + with `c < 0` show ?thesis + by (cases "a \ b") + (auto simp: field_simps emeasure_density positive_integral_distr + positive_integral_cmult borel_measurable_indicator' emeasure_distr) qed -qed +qed simp end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Measure.thy --- a/src/HOL/Probability/Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1452 +0,0 @@ -(* Title: HOL/Probability/Measure.thy - Author: Lawrence C Paulson - Author: Johannes Hölzl, TU München - Author: Armin Heller, TU München -*) - -header {* Properties about measure spaces *} - -theory Measure - 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 algebra_measure_update[simp]: - "algebra (M'\measure := m\) \ algebra M'" - unfolding algebra_iff_Un 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 measurable_cancel_measure[simp]: - "measurable M1 (M2\measure := m2\) = measurable M1 M2" - "measurable (M2\measure := m1\) M1 = measurable M2 M1" - unfolding measurable_def by auto - -lemma inj_on_image_eq_iff: - assumes "inj_on f S" - assumes "A \ S" "B \ S" - shows "(f ` A = f ` B) \ (A = B)" -proof - - have "inj_on f (A \ B)" - using assms by (auto intro: subset_inj_on) - from inj_on_Un_image_eq_iff[OF this] - show ?thesis . -qed - -lemma image_vimage_inter_eq: - assumes "f ` S = T" "X \ T" - shows "f ` (f -` X \ S) = X" -proof (intro antisym) - have "f ` (f -` X \ S) \ f ` (f -` X)" by auto - also have "\ = X \ range f" by simp - also have "\ = X" using assms by auto - finally show "f ` (f -` X \ S) \ X" by auto -next - show "X \ f ` (f -` X \ S)" - proof - fix x assume "x \ X" - then have "x \ T" using `X \ T` by auto - then obtain y where "x = f y" "y \ S" - using assms by auto - then have "{y} \ f -` X \ S" using `x \ X` by auto - moreover have "x \ f ` {y}" using `x = f y` by auto - ultimately show "x \ f ` (f -` X \ S)" by auto - qed -qed - -text {* - This formalisation of measure theory is based on the work of Hurd/Coble wand - was later translated by Lawrence Paulson to Isabelle/HOL. Later it was - modified to use the positive infinite reals and to prove the uniqueness of - cut stable measures. -*} - -section {* Equations for the measure function @{text \} *} - -lemma (in measure_space) measure_countably_additive: - assumes "range A \ sets M" "disjoint_family A" - shows "(\i. \ (A i)) = \ (\i. A i)" -proof - - have "(\ i. A i) \ sets M" using assms(1) by (rule countable_UN) - 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 \ 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 "positive N (measure N)" using assms by (auto simp: positive_def) - 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 "(\n. measure N (A n)) = measure N (UNION UNIV A)" - unfolding assms by simp - qed - qed -qed - -lemma (in measure_space) additive: "additive M \" - using ca by (auto intro!: countably_additive_additive simp: positive_def) - -lemma (in measure_space) measure_additive: - "a \ sets M \ b \ sets M \ a \ b = {} - \ \ a + \ b = \ (a \ b)" - by (metis additiveD additive) - -lemma (in measure_space) measure_mono: - assumes "a \ b" "a \ sets M" "b \ sets M" - shows "\ a \ \ b" -proof - - have "b = a \ (b - a)" using assms by auto - moreover have "{} = a \ (b - a)" by auto - ultimately have "\ b = \ a + \ (b - a)" - using measure_additive[of a "b - a"] Diff[of b a] assms by auto - moreover have "\ a + 0 \ \ a + \ (b - a)" using assms by (intro add_mono) auto - ultimately show "\ a \ \ b" by auto -qed - -lemma (in measure_space) measure_top: - "A \ sets M \ \ A \ \ (space M)" - using sets_into_space[of A] by (auto intro!: measure_mono) - -lemma (in measure_space) measure_compl: - assumes s: "s \ sets M" and fin: "\ s \ \" - shows "\ (space M - s) = \ (space M) - \ s" -proof - - have s_less_space: "\ s \ \ (space M)" - using s by (auto intro!: measure_mono sets_into_space) - from s have "0 \ \ s" by auto - have "\ (space M) = \ (s \ (space M - s))" using s - by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) - also have "... = \ s + \ (space M - s)" - by (rule additiveD [OF additive]) (auto simp add: s) - finally have "\ (space M) = \ s + \ (space M - s)" . - then show ?thesis - using fin `0 \ \ s` - unfolding ereal_eq_minus_iff by (auto simp: ac_simps) -qed - -lemma (in measure_space) measure_Diff: - assumes finite: "\ B \ \" - and measurable: "A \ sets M" "B \ sets M" "B \ A" - shows "\ (A - B) = \ A - \ B" -proof - - have "0 \ \ B" using assms by auto - have "(A - B) \ B = A" using `B \ A` by auto - then have "\ A = \ ((A - B) \ B)" by simp - also have "\ = \ (A - B) + \ B" - using measurable by (subst measure_additive[symmetric]) auto - finally show "\ (A - B) = \ A - \ B" - unfolding ereal_eq_minus_iff - using finite `0 \ \ B` by auto -qed - -lemma (in measure_space) measure_countable_increasing: - assumes A: "range A \ sets M" - and A0: "A 0 = {}" - and ASuc: "\n. A n \ A (Suc n)" - shows "(SUP n. \ (A n)) = \ (\i. A i)" -proof - - { fix n - have "\ (A n) = (\i (A (Suc i) - A i))" - proof (induct n) - case 0 thus ?case by (auto simp add: A0) - next - case (Suc m) - have "A (Suc m) = A m \ (A (Suc m) - A m)" - by (metis ASuc Un_Diff_cancel Un_absorb1) - hence "\ (A (Suc m)) = - \ (A m) + \ (A (Suc m) - A m)" - by (subst measure_additive) - (auto simp add: measure_additive range_subsetD [OF A]) - with Suc show ?case - by simp - qed } - note Meq = this - have Aeq: "(\i. A (Suc i) - A i) = (\i. A i)" - proof (rule UN_finite2_eq [where k=1], simp) - fix i - show "(\i\{0..i\{0..i. A (Suc i) - A i \ sets M" - by (metis A Diff range_subsetD) - have A2: "(\i. A (Suc i) - A i) \ sets M" - by (blast intro: range_subsetD [OF A]) - have "(SUP n. \i (A (Suc i) - A i)) = (\i. \ (A (Suc i) - A i))" - using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric]) - also have "\ = \ (\i. A (Suc i) - A i)" - by (rule measure_countably_additive) - (auto simp add: disjoint_family_Suc ASuc A1 A2) - also have "... = \ (\i. A i)" - by (simp add: Aeq) - finally have "(SUP n. \i (A (Suc i) - A i)) = \ (\i. A i)" . - then show ?thesis by (auto simp add: Meq) -qed - -lemma (in measure_space) continuity_from_below: - assumes A: "range A \ sets M" and "incseq A" - shows "(SUP n. \ (A n)) = \ (\i. A i)" -proof - - have *: "(SUP n. \ (nat_case {} A (Suc n))) = (SUP n. \ (nat_case {} A n))" - using A by (auto intro!: SUPR_eq exI split: nat.split) - have ueq: "(\i. nat_case {} A i) = (\i. A i)" - by (auto simp add: split: nat.splits) - have meq: "\n. \ (A n) = (\ \ nat_case {} A) (Suc n)" - by simp - have "(SUP n. \ (nat_case {} A n)) = \ (\i. nat_case {} A i)" - using range_subsetD[OF A] incseq_SucD[OF `incseq A`] - by (force split: nat.splits intro!: measure_countable_increasing) - also have "\ (\i. nat_case {} A i) = \ (\i. A i)" - by (simp add: ueq) - finally have "(SUP n. \ (nat_case {} A n)) = \ (\i. A i)" . - thus ?thesis unfolding meq * comp_def . -qed - -lemma (in measure_space) measure_incseq: - assumes "range B \ sets M" "incseq B" - shows "incseq (\i. \ (B i))" - using assms by (auto simp: incseq_def intro!: measure_mono) - -lemma (in measure_space) continuity_from_below_Lim: - assumes A: "range A \ sets M" "incseq A" - shows "(\i. (\ (A i))) ----> \ (\i. A i)" - using LIMSEQ_ereal_SUPR[OF measure_incseq, OF A] - continuity_from_below[OF A] by simp - -lemma (in measure_space) measure_decseq: - assumes "range B \ sets M" "decseq B" - shows "decseq (\i. \ (B i))" - using assms by (auto simp: decseq_def intro!: measure_mono) - -lemma (in measure_space) continuity_from_above: - assumes A: "range A \ sets M" and "decseq A" - and finite: "\i. \ (A i) \ \" - shows "(INF n. \ (A n)) = \ (\i. A i)" -proof - - have le_MI: "\ (\i. A i) \ \ (A 0)" - using A by (auto intro!: measure_mono) - hence *: "\ (\i. A i) \ \" using finite[of 0] by auto - - have A0: "0 \ \ (A 0)" using A by auto - - have "\ (A 0) - (INF n. \ (A n)) = \ (A 0) + (SUP n. - \ (A n))" - by (simp add: ereal_SUPR_uminus minus_ereal_def) - also have "\ = (SUP n. \ (A 0) - \ (A n))" - unfolding minus_ereal_def using A0 assms - by (subst SUPR_ereal_add) (auto simp add: measure_decseq) - also have "\ = (SUP n. \ (A 0 - A n))" - using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto - also have "\ = \ (\i. A 0 - A i)" - proof (rule continuity_from_below) - show "range (\n. A 0 - A n) \ sets M" - using A by auto - show "incseq (\n. A 0 - A n)" - using `decseq A` by (auto simp add: incseq_def decseq_def) - qed - also have "\ = \ (A 0) - \ (\i. A i)" - using A finite * by (simp, subst measure_Diff) auto - finally show ?thesis - unfolding ereal_minus_eq_minus_iff using finite A0 by auto -qed - -lemma (in measure_space) measure_insert: - assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" - shows "\ (insert x A) = \ {x} + \ A" -proof - - have "{x} \ A = {}" using `x \ A` by auto - from measure_additive[OF sets this] show ?thesis by simp -qed - -lemma (in measure_space) measure_setsum: - assumes "finite S" and "\i. i \ S \ A i \ sets M" - assumes disj: "disjoint_family_on A S" - shows "(\i\S. \ (A i)) = \ (\i\S. A i)" -using assms proof induct - case (insert i S) - then have "(\i\S. \ (A i)) = \ (\a\S. A a)" - by (auto intro: disjoint_family_on_mono) - moreover have "A i \ (\a\S. A a) = {}" - using `disjoint_family_on A (insert i S)` `i \ S` - by (auto simp: disjoint_family_on_def) - ultimately show ?case using insert - by (auto simp: measure_additive finite_UN) -qed simp - -lemma (in measure_space) measure_finite_singleton: - assumes "finite S" "\x. x \ S \ {x} \ sets M" - shows "\ S = (\x\S. \ {x})" - using measure_setsum[of S "\x. {x}", OF assms] - by (auto simp: disjoint_family_on_def) - -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]: "positive M (measure M)" 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" - and disj: "disjoint_family A" - and UnA: "(\i. A i) \ sets M" - def I \ "{i. A i \ {}}" - have "Union (A ` I) \ space M" using A - by auto (metis range_subsetD subsetD sets_into_space) - hence "finite (A ` I)" - by (metis finite_UnionD finite_subset fin) - moreover have "inj_on A I" using disj - by (auto simp add: I_def disjoint_family_on_def inj_on_def) - ultimately have finI: "finite I" - by (metis finite_imageD) - hence "\N. \m\N. A m = {}" - proof (cases "I = {}") - case True thus ?thesis by (simp add: I_def) - next - case False - hence "\i\I. i < Suc(Max I)" - by (simp add: Max_less_iff [symmetric] finI) - hence "\m \ Suc(Max I). A m = {}" - by (simp add: I_def) (metis less_le_not_le) - thus ?thesis - by blast - qed - then obtain N where N: "\m\N. A m = {}" by blast - then have "\m\N. measure M (A m) = 0" using pos[unfolded positive_def] by simp - then have "(\n. measure M (A n)) = (\mi (\ x i (\ x sets M" using A - by force - show "(\i sets M" - proof (induct n) - case 0 thus ?case by simp - next - case (Suc n) thus ?case using A - by (simp add: lessThan_Suc Un range_subsetD) - qed - qed - thus ?case using Suc - by (simp add: lessThan_Suc) - qed - 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 "(\n. measure M (A n)) = measure M (\i. A i)" . - qed - qed -qed - -lemma (in measure_space) measure_setsum_split: - assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" - assumes "(\i\S. B i) = space M" - assumes "disjoint_family_on B S" - shows "\ A = (\i\S. \ (A \ (B i)))" -proof - - have *: "\ A = \ (\i\S. A \ B i)" - using assms by auto - show ?thesis unfolding * - proof (rule measure_setsum[symmetric]) - show "disjoint_family_on (\i. A \ B i) S" - using `disjoint_family_on B S` - unfolding disjoint_family_on_def by auto - qed (insert assms, auto) -qed - -lemma (in measure_space) measure_subadditive: - assumes measurable: "A \ sets M" "B \ sets M" - shows "\ (A \ B) \ \ A + \ B" -proof - - from measure_additive[of A "B - A"] - have "\ (A \ B) = \ A + \ (B - A)" - using assms by (simp add: Diff) - also have "\ \ \ A + \ B" - using assms by (auto intro!: add_left_mono measure_mono) - finally show ?thesis . -qed - -lemma (in measure_space) measure_subadditive_finite: - assumes "finite I" "\i. i\I \ A i \ sets M" - shows "\ (\i\I. A i) \ (\i\I. \ (A i))" -using assms proof induct - case (insert i I) - then have "\ (\i\insert i I. A i) = \ (A i \ (\i\I. A i))" - by simp - also have "\ \ \ (A i) + \ (\i\I. A i)" - using insert by (intro measure_subadditive finite_UN) auto - also have "\ \ \ (A i) + (\i\I. \ (A i))" - using insert by (intro add_mono) auto - also have "\ = (\i\insert i I. \ (A i))" - using insert by auto - finally show ?case . -qed simp - -lemma (in measure_space) measure_eq_0: - assumes "N \ sets M" and "\ N = 0" and "K \ N" and "K \ sets M" - shows "\ K = 0" - using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto - -lemma (in measure_space) measure_finitely_subadditive: - assumes "finite I" "A ` I \ sets M" - shows "\ (\i\I. A i) \ (\i\I. \ (A i))" -using assms proof induct - case (insert i I) - then have "(\i\I. A i) \ sets M" by auto - then have "\ (\i\insert i I. A i) \ \ (A i) + \ (\i\I. A i)" - using insert by (simp add: measure_subadditive) - also have "\ \ (\i\insert i I. \ (A i))" - using insert by (auto intro!: add_left_mono) - finally show ?case . -qed simp - -lemma (in measure_space) measure_countably_subadditive: - assumes "range f \ sets M" - shows "\ (\i. f i) \ (\i. \ (f i))" -proof - - have "\ (\i. f i) = \ (\i. disjointed f i)" - unfolding UN_disjointed_eq .. - also have "\ = (\i. \ (disjointed f i))" - using range_disjointed_sets[OF assms] measure_countably_additive - by (simp add: disjoint_family_disjointed comp_def) - also have "\ \ (\i. \ (f i))" - using range_disjointed_sets[OF assms] assms - by (auto intro!: suminf_le_pos measure_mono disjointed_subset) - finally show ?thesis . -qed - -lemma (in measure_space) measure_UN_eq_0: - assumes "\i::nat. \ (N i) = 0" and "range N \ sets M" - shows "\ (\ i. N i) = 0" -proof - - have "0 \ \ (\ i. N i)" using assms by auto - moreover have "\ (\ i. N i) \ 0" - using measure_countably_subadditive[OF assms(2)] assms(1) by simp - ultimately show ?thesis by simp -qed - -lemma (in measure_space) measure_inter_full_set: - assumes "S \ sets M" "T \ sets M" and fin: "\ (T - S) \ \" - assumes T: "\ T = \ (space M)" - shows "\ (S \ T) = \ S" -proof (rule antisym) - show " \ (S \ T) \ \ S" - using assms by (auto intro!: measure_mono) - - have pos: "0 \ \ (T - S)" using assms by auto - show "\ S \ \ (S \ T)" - proof (rule ccontr) - assume contr: "\ ?thesis" - have "\ (space M) = \ ((T - S) \ (S \ T))" - unfolding T[symmetric] by (auto intro!: arg_cong[where f="\"]) - also have "\ \ \ (T - S) + \ (S \ T)" - using assms by (auto intro!: measure_subadditive) - also have "\ < \ (T - S) + \ S" - using fin contr pos by (intro ereal_less_add) auto - also have "\ = \ (T \ S)" - using assms by (subst measure_additive) auto - also have "\ \ \ (space M)" - using assms sets_into_space by (auto intro!: measure_mono) - finally show False .. - qed -qed - -lemma measure_unique_Int_stable: - fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \ 'a set" - assumes "Int_stable E" - and A: "range A \ sets E" "incseq A" "(\i. A i) = 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 (sigma E)" - shows "\ X = \ X" -proof - - 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 (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 (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` by auto - then show "\ (F \ space E) = \ (F \ space E)" - using `F \ sets E` eq by auto - next - 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 (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 (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 (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] - N.measure_countably_additive[symmetric] - simp del: UN_simps) - qed - 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 (sigma E) \ \ (F \ D) = \ (F \ D)" - by (subst (asm) *) auto } - note * = this - let ?A = "\i. A i \ X" - have A': "range ?A \ sets (sigma E)" "incseq ?A" - using A(1,2) `X \ sets (sigma E)` by (auto simp: incseq_def) - { fix i have "\ (?A i) = \ (?A i)" - using *[of "A i" X] `X \ sets (sigma E)` A finite by auto } - with M.continuity_from_below[OF A'] N.continuity_from_below[OF A'] - show ?thesis using A(3) `X \ sets (sigma E)` by auto -qed - -section "@{text \}-null sets" - -abbreviation (in measure_space) "null_sets \ {N\sets M. \ N = 0}" - -sublocale measure_space \ nullsets!: ring_of_sets "\ space = space M, sets = null_sets \" - where "space \ space = space M, sets = null_sets \ = space M" - and "sets \ space = space M, sets = null_sets \ = null_sets" -proof - - { fix A B assume sets: "A \ sets M" "B \ sets M" - moreover then have "\ (A \ B) \ \ A + \ B" "\ (A - B) \ \ A" - by (auto intro!: measure_subadditive measure_mono) - moreover assume "\ B = 0" "\ A = 0" - ultimately have "\ (A - B) = 0" "\ (A \ B) = 0" - by (auto intro!: antisym) } - note null = this - show "ring_of_sets \ space = space M, sets = null_sets \" - by default (insert sets_into_space null, auto) -qed simp_all - -lemma UN_from_nat: "(\i. N i) = (\i. N (Countable.from_nat i))" -proof - - have "(\i. N i) = (\i. (N \ Countable.from_nat) i)" - unfolding SUP_def image_compose - unfolding surj_from_nat .. - then show ?thesis by simp -qed - -lemma (in measure_space) null_sets_UN[intro]: - assumes "\i::'i::countable. N i \ null_sets" - shows "(\i. N i) \ null_sets" -proof (intro conjI CollectI) - show "(\i. N i) \ sets M" using assms by auto - then have "0 \ \ (\i. N i)" by simp - moreover have "\ (\i. N i) \ (\n. \ (N (Countable.from_nat n)))" - unfolding UN_from_nat[of N] - using assms by (intro measure_countably_subadditive) auto - ultimately show "\ (\i. N i) = 0" using assms by auto -qed - -lemma (in measure_space) null_set_Int1: - assumes "B \ null_sets" "A \ sets M" shows "A \ B \ null_sets" -using assms proof (intro CollectI conjI) - show "\ (A \ B) = 0" using assms by (intro measure_eq_0[of B "A \ B"]) auto -qed auto - -lemma (in measure_space) null_set_Int2: - assumes "B \ null_sets" "A \ sets M" shows "B \ A \ null_sets" - using assms by (subst Int_commute) (rule null_set_Int1) - -lemma (in measure_space) measure_Diff_null_set: - assumes "B \ null_sets" "A \ sets M" - shows "\ (A - B) = \ A" -proof - - have *: "A - B = (A - (A \ B))" by auto - have "A \ B \ null_sets" using assms by (rule null_set_Int1) - then show ?thesis - unfolding * using assms - by (subst measure_Diff) auto -qed - -lemma (in measure_space) null_set_Diff: - assumes "B \ null_sets" "A \ sets M" shows "B - A \ null_sets" -using assms proof (intro CollectI conjI) - show "\ (B - A) = 0" using assms by (intro measure_eq_0[of B "B - A"]) auto -qed auto - -lemma (in measure_space) measure_Un_null_set: - assumes "A \ sets M" "B \ null_sets" - shows "\ (A \ B) = \ A" -proof - - have *: "A \ B = A \ (B - A)" by auto - have "B - A \ null_sets" using assms(2,1) by (rule null_set_Diff) - then show ?thesis - unfolding * using assms - by (subst measure_additive[symmetric]) auto -qed - -section "Formalise almost everywhere" - -definition (in measure_space) - almost_everywhere :: "('a \ bool) \ bool" (binder "AE " 10) where - "almost_everywhere P \ (\N\null_sets. { x \ space M. \ P x } \ N)" - -syntax - "_almost_everywhere" :: "pttrn \ ('a, 'b) measure_space_scheme \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) - -translations - "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)" - -lemma (in measure_space) AE_cong_measure: - assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" - shows "(AE x in N. P x) \ (AE x. P x)" -proof - - interpret N: measure_space N - by (rule measure_space_cong) fact+ - show ?thesis - unfolding N.almost_everywhere_def almost_everywhere_def - by (auto simp: assms) -qed - -lemma (in measure_space) AE_I': - "N \ null_sets \ {x\space M. \ P x} \ N \ (AE x. P x)" - unfolding almost_everywhere_def by auto - -lemma (in measure_space) AE_iff_null_set: - assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") - shows "(AE x. P x) \ {x\space M. \ P x} \ null_sets" -proof - assume "AE x. P x" then obtain N where N: "N \ sets M" "?P \ N" "\ N = 0" - unfolding almost_everywhere_def by auto - have "0 \ \ ?P" using assms by simp - moreover have "\ ?P \ \ N" - using assms N(1,2) by (auto intro: measure_mono) - ultimately have "\ ?P = 0" unfolding `\ N = 0` by auto - then show "?P \ null_sets" using assms by simp -next - assume "?P \ null_sets" with assms show "AE x. P x" by (auto intro: AE_I') -qed - -lemma (in measure_space) AE_iff_measurable: - "N \ sets M \ {x\space M. \ P x} = N \ (AE x. P x) \ \ N = 0" - using AE_iff_null_set[of P] by simp - -lemma (in measure_space) AE_True[intro, simp]: "AE x. True" - unfolding almost_everywhere_def by auto - -lemma (in measure_space) AE_E[consumes 1]: - assumes "AE x. P x" - obtains N where "{x \ space M. \ P x} \ N" "\ N = 0" "N \ sets M" - using assms unfolding almost_everywhere_def by auto - -lemma (in measure_space) AE_E2: - assumes "AE x. P x" "{x\space M. P x} \ sets M" - shows "\ {x\space M. \ P x} = 0" (is "\ ?P = 0") -proof - - have "{x\space M. \ P x} = space M - {x\space M. P x}" - by auto - with AE_iff_null_set[of P] assms show ?thesis by auto -qed - -lemma (in measure_space) AE_I: - assumes "{x \ space M. \ P x} \ N" "\ N = 0" "N \ sets M" - shows "AE x. P x" - using assms unfolding almost_everywhere_def by auto - -lemma (in measure_space) AE_mp[elim!]: - assumes AE_P: "AE x. P x" and AE_imp: "AE x. P x \ Q x" - shows "AE x. Q x" -proof - - from AE_P obtain A where P: "{x\space M. \ P x} \ A" - and A: "A \ sets M" "\ A = 0" - by (auto elim!: AE_E) - - from AE_imp obtain B where imp: "{x\space M. P x \ \ Q x} \ B" - and B: "B \ sets M" "\ B = 0" - by (auto elim!: AE_E) - - show ?thesis - proof (intro AE_I) - have "0 \ \ (A \ B)" using A B by auto - moreover have "\ (A \ B) \ 0" - using measure_subadditive[of A B] A B by auto - ultimately show "A \ B \ sets M" "\ (A \ B) = 0" using A B by auto - show "{x\space M. \ Q x} \ A \ B" - using P imp by auto - qed -qed - -lemma (in measure_space) - shows AE_iffI: "AE x. P x \ AE x. P x \ Q x \ AE x. Q x" - and AE_disjI1: "AE x. P x \ AE x. P x \ Q x" - and AE_disjI2: "AE x. Q x \ AE x. P x \ Q x" - and AE_conjI: "AE x. P x \ AE x. Q x \ AE x. P x \ Q x" - and AE_conj_iff[simp]: "(AE x. P x \ Q x) \ (AE x. P x) \ (AE x. Q x)" - by auto - -lemma (in measure_space) AE_measure: - assumes AE: "AE x. P x" and sets: "{x\space M. P x} \ sets M" - shows "\ {x\space M. P x} = \ (space M)" -proof - - from AE_E[OF AE] guess N . note N = this - with sets have "\ (space M) \ \ ({x\space M. P x} \ N)" - by (intro measure_mono) auto - also have "\ \ \ {x\space M. P x} + \ N" - using sets N by (intro measure_subadditive) auto - also have "\ = \ {x\space M. P x}" using N by simp - finally show "\ {x\space M. P x} = \ (space M)" - using measure_top[OF sets] by auto -qed - -lemma (in measure_space) AE_space: "AE x. x \ space M" - by (rule AE_I[where N="{}"]) auto - -lemma (in measure_space) AE_I2[simp, intro]: - "(\x. x \ space M \ P x) \ AE x. P x" - using AE_space by auto - -lemma (in measure_space) AE_Ball_mp: - "\x\space M. P x \ AE x. P x \ Q x \ AE x. Q x" - by auto - -lemma (in measure_space) AE_cong[cong]: - "(\x. x \ space M \ P x \ Q x) \ (AE x. P x) \ (AE x. Q x)" - by auto - -lemma (in measure_space) AE_all_countable: - "(AE x. \i. P i x) \ (\i::'i::countable. AE x. P i x)" -proof - assume "\i. AE x. P i x" - from this[unfolded almost_everywhere_def Bex_def, THEN choice] - obtain N where N: "\i. N i \ null_sets" "\i. {x\space M. \ P i x} \ N i" by auto - have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto - also have "\ \ (\i. N i)" using N by auto - finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . - moreover from N have "(\i. N i) \ null_sets" - by (intro null_sets_UN) auto - ultimately show "AE x. \i. P i x" - unfolding almost_everywhere_def by auto -qed auto - -lemma (in measure_space) AE_finite_all: - assumes f: "finite S" shows "(AE x. \i\S. P i x) \ (\i\S. AE x. P i x)" - using f by induct auto - -lemma (in measure_space) restricted_measure_space: - assumes "S \ sets M" - 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 "positive ?r (measure ?r)" using `S \ sets M` by (auto simp: positive_def) - - 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 "(\n. measure ?r (A n)) = measure ?r (\i. A i)" - using measure_countably_additive by simp - qed -qed - -lemma (in measure_space) AE_restricted: - assumes "A \ sets M" - shows "(AE x in restricted_space A. P x) \ (AE x. x \ A \ P x)" -proof - - interpret R: measure_space "restricted_space A" - by (rule restricted_measure_space[OF `A \ sets M`]) - show ?thesis - proof - assume "AE x in restricted_space A. P x" - from this[THEN R.AE_E] guess N' . - then obtain N where "{x \ A. \ P x} \ A \ N" "\ (A \ N) = 0" "N \ sets M" - by auto - moreover then have "{x \ space M. \ (x \ A \ P x)} \ A \ N" - using `A \ sets M` sets_into_space by auto - ultimately show "AE x. x \ A \ P x" - using `A \ sets M` by (auto intro!: AE_I[where N="A \ N"]) - next - assume "AE x. x \ A \ P x" - from this[THEN AE_E] guess N . - then show "AE x in restricted_space A. P x" - using null_set_Int1[OF _ `A \ sets M`] `A \ sets M`[THEN sets_into_space] - by (auto intro!: R.AE_I[where N="A \ N"] simp: subset_eq) - qed -qed - -lemma (in measure_space) measure_space_subalgebra: - assumes "sigma_algebra N" and "sets N \ sets M" "space N = space M" - 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 (measure N)" - by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq) - show "positive N (measure_space.measure N)" - using assms(2) by (auto simp add: positive_def) - qed -qed - -lemma (in measure_space) AE_subalgebra: - assumes ae: "AE x in N. P x" - and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" - and sa: "sigma_algebra N" - shows "AE x. P x" -proof - - interpret N: measure_space N using measure_space_subalgebra[OF sa N] . - from ae[THEN N.AE_E] guess N . - with N show ?thesis unfolding almost_everywhere_def by auto -qed - -section "@{text \}-finite Measures" - -locale sigma_finite_measure = measure_space + - assumes sigma_finite: "\A::nat \ 'a set. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" - -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") - unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def -proof safe - 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. measure ?r (A i) \ \)" - proof (safe intro!: exI[of _ "\i. A i \ S"] del: notI) - fix i - show "A i \ S \ sets ?r" - using `range A \ sets M` `S \ sets M` by auto - next - fix x i assume "x \ S" thus "x \ space ?r" by simp - next - fix x assume "x \ space ?r" thus "x \ (\i. A i \ S)" - using `(\i. A i) = space M` `S \ sets M` by auto - next - fix i - have "\ (A i \ S) \ \ (A i)" - using `range A \ sets M` `S \ sets M` by (auto intro!: measure_mono) - then show "measure ?r (A i \ S) \ \" using `\ (A i) \ \` by auto - qed -qed - -lemma (in sigma_finite_measure) sigma_finite_measure_cong: - 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 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. measure M' (A i) \ \" using cong by auto - show ?thesis - apply default - using A fin cong by auto -qed - -lemma (in sigma_finite_measure) disjoint_sigma_finite: - "\A::nat\'a set. range A \ sets M \ (\i. A i) = space M \ - (\i. \ (A i) \ \) \ disjoint_family A" -proof - - obtain A :: "nat \ 'a set" where - range: "range A \ sets M" and - space: "(\i. A i) = space M" and - measure: "\i. \ (A i) \ \" - using sigma_finite by auto - note range' = range_disjointed_sets[OF range] range - { fix i - have "\ (disjointed A i) \ \ (A i)" - using range' disjointed_subset[of A i] by (auto intro!: measure_mono) - then have "\ (disjointed A i) \ \" - using measure[of i] by auto } - with disjoint_family_disjointed UN_disjointed_eq[of A] space range' - show ?thesis by (auto intro!: exI[of _ "disjointed A"]) -qed - -lemma (in sigma_finite_measure) sigma_finite_up: - "\F. range F \ sets M \ incseq F \ (\i. F i) = space M \ (\i. \ (F i) \ \)" -proof - - obtain F :: "nat \ 'a set" where - F: "range F \ sets M" "(\i. F i) = space M" "\i. \ (F i) \ \" - using sigma_finite by auto - then show ?thesis - proof (intro exI[of _ "\n. \i\n. F i"] conjI allI) - from F have "\x. x \ space M \ \i. x \ F i" by auto - then show "(\n. \ i\n. F i) = space M" - using F by fastforce - next - fix n - have "\ (\ i\n. F i) \ (\i\n. \ (F i))" using F - by (auto intro!: measure_finitely_subadditive) - also have "\ < \" - using F by (auto simp: setsum_Pinfty) - finally show "\ (\ i\n. F i) \ \" by simp - qed (force simp: incseq_def)+ -qed - -section {* Measure preserving *} - -definition "measure_preserving A B = - {f \ measurable A B. (\y \ sets B. measure B y = measure A (f -` y \ space A))}" - -lemma measure_preservingI[intro?]: - assumes "f \ measurable A B" - and "\y. y \ sets B \ measure A (f -` y \ space A) = measure B y" - shows "f \ measure_preserving A B" - unfolding measure_preserving_def using assms by auto - -lemma (in measure_space) measure_space_vimage: - fixes M' :: "('c, 'd) measure_space_scheme" - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - shows "measure_space M'" -proof - - interpret M': sigma_algebra M' by fact - show ?thesis - proof - show "positive M' (measure M')" using T - by (auto simp: measure_preserving_def positive_def measurable_sets) - - 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 by (auto simp: measurable_def measure_preserving_def) - moreover have "(\i. T -` A i \ space M) \ sets M" - 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 "(\i. measure M' (A i)) = measure M' (\i. A i)" - using measure_countably_additive[OF _ **] A T - by (auto simp: comp_def vimage_UN measure_preserving_def) - qed - qed -qed - -lemma (in measure_space) almost_everywhere_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" - and AE: "measure_space.almost_everywhere M' P" - shows "AE x. P (T x)" -proof - - interpret M': measure_space M' using T by (rule measure_space_vimage) - from AE[THEN M'.AE_E] guess N . - then show ?thesis - unfolding almost_everywhere_def M'.almost_everywhere_def - using T(2) unfolding measurable_def measure_preserving_def - by (intro bexI[of _ "T -` N \ space M"]) auto -qed - -lemma measure_unique_Int_stable_vimage: - fixes A :: "nat \ 'a set" - assumes E: "Int_stable E" - and A: "range A \ sets E" "incseq A" "(\i. A i) = space E" "\i. measure M (A i) \ \" - and N: "measure_space N" "T \ measurable N M" - and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M" - and eq: "\X. X \ sets E \ measure M X = measure N (T -` X \ space N)" - assumes X: "X \ sets (sigma E)" - shows "measure M X = measure N (T -` X \ space N)" -proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X]) - interpret M: measure_space M by fact - interpret N: measure_space N by fact - let ?T = "\X. T -` X \ space N" - show "measure_space \space = space E, sets = sets (sigma E), measure = measure M\" - by (rule M.measure_space_cong) (auto simp: M) - show "measure_space \space = space E, sets = sets (sigma E), measure = \X. measure N (?T X)\" (is "measure_space ?E") - proof (rule N.measure_space_vimage) - show "sigma_algebra ?E" - by (rule M.sigma_algebra_cong) (auto simp: M) - show "T \ measure_preserving N ?E" - using `T \ measurable N M` by (auto simp: M measurable_def measure_preserving_def) - qed - show "\i. M.\ (A i) \ \" by fact -qed - -lemma (in measure_space) measure_preserving_Int_stable: - fixes A :: "nat \ 'a set" - assumes E: "Int_stable E" "range A \ sets E" "incseq A" "(\i. A i) = space E" "\i. measure E (A i) \ \" - and N: "measure_space (sigma E)" - and T: "T \ measure_preserving M E" - shows "T \ measure_preserving M (sigma E)" -proof - interpret E: measure_space "sigma E" by fact - show "T \ measurable M (sigma E)" - using T E.sets_into_space - by (intro measurable_sigma) (auto simp: measure_preserving_def measurable_def) - fix X assume "X \ sets (sigma E)" - show "\ (T -` X \ space M) = E.\ X" - proof (rule measure_unique_Int_stable_vimage[symmetric]) - show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)" - "\i. E.\ (A i) \ \" using E by auto - show "measure_space M" by default - next - fix X assume "X \ sets E" then show "E.\ X = \ (T -` X \ space M)" - using T unfolding measure_preserving_def by auto - qed fact+ -qed - -section "Real measure values" - -lemma (in measure_space) real_measure_Union: - assumes finite: "\ A \ \" "\ B \ \" - and measurable: "A \ sets M" "B \ sets M" "A \ B = {}" - shows "real (\ (A \ B)) = real (\ A) + real (\ B)" - unfolding measure_additive[symmetric, OF measurable] - using measurable(1,2)[THEN positive_measure] - using finite by (cases rule: ereal2_cases[of "\ A" "\ B"]) auto - -lemma (in measure_space) real_measure_finite_Union: - assumes measurable: - "finite S" "\i. i \ S \ A i \ sets M" "disjoint_family_on A S" - assumes finite: "\i. i \ S \ \ (A i) \ \" - shows "real (\ (\i\S. A i)) = (\i\S. real (\ (A i)))" - using finite measurable(2)[THEN positive_measure] - by (force intro!: setsum_real_of_ereal[symmetric] - simp: measure_setsum[OF measurable, symmetric]) - -lemma (in measure_space) real_measure_Diff: - assumes finite: "\ A \ \" - and measurable: "A \ sets M" "B \ sets M" "B \ A" - shows "real (\ (A - B)) = real (\ A) - real (\ B)" -proof - - have "\ (A - B) \ \ A" "\ B \ \ A" - using measurable by (auto intro!: measure_mono) - hence "real (\ ((A - B) \ B)) = real (\ (A - B)) + real (\ B)" - using measurable finite by (rule_tac real_measure_Union) auto - thus ?thesis using `B \ A` by (auto simp: Un_absorb2) -qed - -lemma (in measure_space) real_measure_UNION: - assumes measurable: "range A \ sets M" "disjoint_family A" - assumes finite: "\ (\i. A i) \ \" - shows "(\i. real (\ (A i))) sums (real (\ (\i. A i)))" -proof - - have "\i. 0 \ \ (A i)" using measurable by auto - with summable_sums[OF summable_ereal_pos, of "\i. \ (A i)"] - measure_countably_additive[OF measurable] - have "(\i. \ (A i)) sums (\ (\i. A i))" by simp - moreover - { fix i - have "\ (A i) \ \ (\i. A i)" - using measurable by (auto intro!: measure_mono) - moreover have "0 \ \ (A i)" using measurable by auto - ultimately have "\ (A i) = ereal (real (\ (A i)))" - using finite by (cases "\ (A i)") auto } - moreover - have "0 \ \ (\i. A i)" using measurable by auto - then have "\ (\i. A i) = ereal (real (\ (\i. A i)))" - using finite by (cases "\ (\i. A i)") auto - ultimately show ?thesis - unfolding sums_ereal[symmetric] by simp -qed - -lemma (in measure_space) real_measure_subadditive: - assumes measurable: "A \ sets M" "B \ sets M" - and fin: "\ A \ \" "\ B \ \" - shows "real (\ (A \ B)) \ real (\ A) + real (\ B)" -proof - - have "0 \ \ (A \ B)" using measurable by auto - then show "real (\ (A \ B)) \ real (\ A) + real (\ B)" - using measure_subadditive[OF measurable] fin - by (cases rule: ereal3_cases[of "\ (A \ B)" "\ A" "\ B"]) auto -qed - -lemma (in measure_space) real_measure_setsum_singleton: - assumes S: "finite S" "\x. x \ S \ {x} \ sets M" - and fin: "\x. x \ S \ \ {x} \ \" - shows "real (\ S) = (\x\S. real (\ {x}))" - using measure_finite_singleton[OF S] fin - using positive_measure[OF S(2)] - by (force intro!: setsum_real_of_ereal[symmetric]) - -lemma (in measure_space) real_continuity_from_below: - assumes A: "range A \ sets M" "incseq A" and fin: "\ (\i. A i) \ \" - shows "(\i. real (\ (A i))) ----> real (\ (\i. A i))" -proof - - have "0 \ \ (\i. A i)" using A by auto - then have "ereal (real (\ (\i. A i))) = \ (\i. A i)" - using fin by (auto intro: ereal_real') - then show ?thesis - using continuity_from_below_Lim[OF A] - by (intro lim_real_of_ereal) simp -qed - -lemma (in measure_space) continuity_from_above_Lim: - assumes A: "range A \ sets M" "decseq A" and fin: "\i. \ (A i) \ \" - shows "(\i. (\ (A i))) ----> \ (\i. A i)" - using LIMSEQ_ereal_INFI[OF measure_decseq, OF A] - using continuity_from_above[OF A fin] by simp - -lemma (in measure_space) real_continuity_from_above: - assumes A: "range A \ sets M" "decseq A" and fin: "\i. \ (A i) \ \" - shows "(\n. real (\ (A n))) ----> real (\ (\i. A i))" -proof - - have "0 \ \ (\i. A i)" using A by auto - moreover - have "\ (\i. A i) \ \ (A 0)" - using A by (auto intro!: measure_mono) - ultimately have "ereal (real (\ (\i. A i))) = \ (\i. A i)" - using fin by (auto intro: ereal_real') - then show ?thesis - using continuity_from_above_Lim[OF A fin] - by (intro lim_real_of_ereal) simp -qed - -lemma (in measure_space) real_measure_countably_subadditive: - assumes A: "range A \ sets M" and fin: "(\i. \ (A i)) \ \" - shows "real (\ (\i. A i)) \ (\i. real (\ (A i)))" -proof - - { fix i - have "0 \ \ (A i)" using A by auto - moreover have "\ (A i) \ \" using A by (intro suminf_PInfty[OF _ fin]) auto - ultimately have "\\ (A i)\ \ \" by auto } - moreover have "0 \ \ (\i. A i)" using A by auto - ultimately have "ereal (real (\ (\i. A i))) \ (\i. ereal (real (\ (A i))))" - using measure_countably_subadditive[OF A] by (auto simp: ereal_real) - also have "\ = ereal (\i. real (\ (A i)))" - using A - by (auto intro!: sums_unique[symmetric] sums_ereal[THEN iffD2] summable_sums summable_real_of_ereal fin) - finally show ?thesis by simp -qed - -locale finite_measure = sigma_finite_measure + - assumes finite_measure_of_space: "\ (space M) \ \" - -lemma finite_measureI[Pure.intro!]: - assumes "measure_space M" - assumes *: "measure M (space M) \ \" - shows "finite_measure M" -proof - - interpret measure_space M by fact - show "finite_measure M" - proof - show "measure M (space M) \ \" by fact - show "\A. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" - using * by (auto intro!: exI[of _ "\x. space M"]) - qed -qed - -lemma (in finite_measure) finite_measure[simp, intro]: - assumes "A \ sets M" - shows "\ A \ \" -proof - - from `A \ sets M` have "A \ space M" - using sets_into_space by blast - then have "\ A \ \ (space M)" - using assms top by (rule measure_mono) - then show ?thesis - using finite_measure_of_space by auto -qed - -definition (in finite_measure) - "\' A = (if A \ sets M then real (\ A) else 0)" - -lemma (in finite_measure) finite_measure_eq: "A \ sets M \ \ A = ereal (\' A)" - by (auto simp: \'_def ereal_real) - -lemma (in finite_measure) positive_measure'[simp, intro]: "0 \ \' A" - unfolding \'_def by (auto simp: real_of_ereal_pos) - -lemma (in finite_measure) real_measure: - assumes A: "A \ sets M" shows "\r. 0 \ r \ \ A = ereal r" - using finite_measure[OF A] positive_measure[OF A] by (cases "\ A") auto - -lemma (in finite_measure) bounded_measure: "\' A \ \' (space M)" -proof cases - assume "A \ sets M" - moreover then have "\ A \ \ (space M)" - using sets_into_space by (auto intro!: measure_mono) - ultimately show ?thesis - by (auto simp: \'_def intro!: real_of_ereal_positive_mono) -qed (simp add: \'_def real_of_ereal_pos) - -lemma (in finite_measure) restricted_finite_measure: - assumes "S \ sets M" - shows "finite_measure (restricted_space S)" - (is "finite_measure ?r") -proof - show "measure_space ?r" using restricted_measure_space[OF assms] . - 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)" -proof - show "measure_space (restricted_space S)" - using `S \ sets M` by (rule restricted_measure_space) - show "measure (restricted_space S) (space (restricted_space S)) \ \" - by simp fact -qed - -lemma (in finite_measure) finite_measure_Diff: - assumes sets: "A \ sets M" "B \ sets M" and "B \ A" - shows "\' (A - B) = \' A - \' B" - using sets[THEN finite_measure_eq] - using Diff[OF sets, THEN finite_measure_eq] - using measure_Diff[OF _ assms] by simp - -lemma (in finite_measure) finite_measure_Union: - assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" - shows "\' (A \ B) = \' A + \' B" - using measure_additive[OF assms] - using sets[THEN finite_measure_eq] - using Un[OF sets, THEN finite_measure_eq] - by simp - -lemma (in finite_measure) finite_measure_finite_Union: - assumes S: "finite S" "\i. i \ S \ A i \ sets M" - and dis: "disjoint_family_on A S" - shows "\' (\i\S. A i) = (\i\S. \' (A i))" - using measure_setsum[OF assms] - using finite_UN[of S A, OF S, THEN finite_measure_eq] - using S(2)[THEN finite_measure_eq] - by simp - -lemma (in finite_measure) finite_measure_UNION: - assumes A: "range A \ sets M" "disjoint_family A" - shows "(\i. \' (A i)) sums (\' (\i. A i))" - using real_measure_UNION[OF A] - using countable_UN[OF A(1), THEN finite_measure_eq] - using A(1)[THEN subsetD, THEN finite_measure_eq] - by auto - -lemma (in finite_measure) finite_measure_mono: - assumes B: "B \ sets M" and "A \ B" shows "\' A \ \' B" -proof cases - assume "A \ sets M" - from this[THEN finite_measure_eq] B[THEN finite_measure_eq] - show ?thesis using measure_mono[OF `A \ B` `A \ sets M` `B \ sets M`] by simp -next - assume "A \ sets M" then show ?thesis - using positive_measure'[of B] unfolding \'_def by auto -qed - -lemma (in finite_measure) finite_measure_subadditive: - assumes m: "A \ sets M" "B \ sets M" - shows "\' (A \ B) \ \' A + \' B" - using measure_subadditive[OF m] - using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp - -lemma (in finite_measure) finite_measure_subadditive_finite: - assumes "finite I" "\i. i\I \ A i \ sets M" - shows "\' (\i\I. A i) \ (\i\I. \' (A i))" - using measure_subadditive_finite[OF assms] assms - by (simp add: finite_measure_eq finite_UN) - -lemma (in finite_measure) finite_measure_countably_subadditive: - assumes A: "range A \ sets M" and sum: "summable (\i. \' (A i))" - shows "\' (\i. A i) \ (\i. \' (A i))" -proof - - note A[THEN subsetD, THEN finite_measure_eq, simp] - note countable_UN[OF A, THEN finite_measure_eq, simp] - from `summable (\i. \' (A i))` - have "(\i. ereal (\' (A i))) sums ereal (\i. \' (A i))" - by (simp add: sums_ereal) (rule summable_sums) - from sums_unique[OF this, symmetric] - measure_countably_subadditive[OF A] - show ?thesis by simp -qed - -lemma (in finite_measure) finite_measure_finite_singleton: - assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" - shows "\' S = (\x\S. \' {x})" - using real_measure_setsum_singleton[OF assms] - using *[THEN finite_measure_eq] - using finite_UN[of S "\x. {x}", OF assms, THEN finite_measure_eq] - by simp - -lemma (in finite_measure) finite_continuity_from_below: - assumes A: "range A \ sets M" and "incseq A" - shows "(\i. \' (A i)) ----> \' (\i. A i)" - using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms - using A[THEN subsetD, THEN finite_measure_eq] - using countable_UN[OF A, THEN finite_measure_eq] - by auto - -lemma (in finite_measure) finite_continuity_from_above: - assumes A: "range A \ sets M" and "decseq A" - shows "(\n. \' (A n)) ----> \' (\i. A i)" - using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms - using A[THEN subsetD, THEN finite_measure_eq] - using countable_INT[OF A, THEN finite_measure_eq] - by auto - -lemma (in finite_measure) finite_measure_compl: - assumes S: "S \ sets M" - shows "\' (space M - S) = \' (space M) - \' S" - using measure_compl[OF S, OF finite_measure, OF S] - using S[THEN finite_measure_eq] - using compl_sets[OF S, THEN finite_measure_eq] - using top[THEN finite_measure_eq] - by simp - -lemma (in finite_measure) finite_measure_inter_full_set: - assumes S: "S \ sets M" "T \ sets M" - assumes T: "\' T = \' (space M)" - shows "\' (S \ T) = \' S" - using measure_inter_full_set[OF S finite_measure] - using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq] - using Int[OF S, THEN finite_measure_eq] - using S[THEN finite_measure_eq] top[THEN finite_measure_eq] - by simp - -lemma (in finite_measure) empty_measure'[simp]: "\' {} = 0" - unfolding \'_def by simp - -section "Finite spaces" - -locale finite_measure_space = finite_measure + finite_sigma_algebra - -lemma finite_measure_spaceI[Pure.intro!]: - assumes "finite (space M)" - assumes sets_Pow: "sets M = Pow (space M)" - and space: "measure M (space M) \ \" - and pos: "\x. x \ space M \ 0 \ measure M {x}" - and add: "\A. A \ space M \ measure M A = (\x\A. measure M {x})" - shows "finite_measure_space M" -proof - - interpret finite_sigma_algebra M - proof - show "finite (space M)" by fact - qed (auto simp: sets_Pow) - interpret measure_space M - proof (rule finite_additivity_sufficient) - show "sigma_algebra M" by default - show "finite (space M)" by fact - show "positive M (measure M)" - by (auto simp: add positive_def intro!: setsum_nonneg pos) - show "additive M (measure M)" - using `finite (space M)` - by (auto simp add: additive_def add - intro!: setsum_Un_disjoint dest: finite_subset) - qed - interpret finite_measure M - proof - show "\ (space M) \ \" by fact - qed default - show "finite_measure_space M" - by default -qed - -lemma (in finite_measure_space) sum_over_space: "(\x\space M. \ {x}) = \ (space M)" - using measure_setsum[of "space M" "\i. {i}"] - by (simp add: disjoint_family_on_def finite_space) - -lemma (in finite_measure_space) finite_measure_singleton: - assumes A: "A \ space M" shows "\' A = (\x\A. \' {x})" - using A finite_subset[OF A finite_space] - by (intro finite_measure_finite_singleton) auto - -lemma (in finite_measure_space) finite_measure_subadditive_setsum: - assumes "finite I" - shows "\' (\i\I. A i) \ (\i\I. \' (A i))" -proof cases - assume "(\i\I. A i) \ space M" - then have "\i. i\I \ A i \ sets M" by auto - from finite_measure_subadditive_finite[OF `finite I` this] - show ?thesis by auto -next - assume "\ (\i\I. A i) \ space M" - then have "\' (\i\I. A i) = 0" - by (simp add: \'_def) - also have "0 \ (\i\I. \' (A i))" - by (auto intro!: setsum_nonneg) - finally show ?thesis . -qed - -lemma suminf_cmult_indicator: - fixes f :: "nat \ ereal" - assumes "disjoint_family A" "x \ A i" "\i. 0 \ f i" - shows "(\n. f n * indicator (A n) x) = f i" -proof - - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)" - using `x \ A i` assms unfolding disjoint_family_on_def indicator_def by auto - then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" - from this[of "Suc i"] show "f i \ y" by auto - qed (insert assms, simp) - ultimately show ?thesis using assms - by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def) -qed - -lemma suminf_indicator: - assumes "disjoint_family A" - shows "(\n. indicator (A n) x :: ereal) = indicator (\i. A i) x" -proof cases - assume *: "x \ (\i. A i)" - then obtain i where "x \ A i" by auto - from suminf_cmult_indicator[OF assms(1), OF `x \ A i`, of "\k. 1"] - show ?thesis using * by simp -qed simp - -end \ No newline at end of file diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Measure_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Measure_Space.thy Mon Apr 23 12:14:35 2012 +0200 @@ -0,0 +1,1457 @@ +(* Title: HOL/Probability/Measure_Space.thy + Author: Lawrence C Paulson + Author: Johannes Hölzl, TU München + Author: Armin Heller, TU München +*) + +header {* Measure spaces and their properties *} + +theory Measure_Space +imports + Sigma_Algebra + "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits" +begin + +lemma suminf_eq_setsum: + fixes f :: "nat \ 'a::{comm_monoid_add, t2_space}" + assumes "finite {i. f i \ 0}" (is "finite ?P") + shows "(\i. f i) = (\i | f i \ 0. f i)" +proof cases + assume "?P \ {}" + have [dest!]: "\i. Suc (Max ?P) \ i \ f i = 0" + using `finite ?P` `?P \ {}` by (auto simp: Suc_le_eq) + have "(\i. f i) = (\i = (\i | f i \ 0. f i)" + using `finite ?P` `?P \ {}` + by (intro setsum_mono_zero_right) (auto simp: less_Suc_eq_le) + finally show ?thesis . +qed simp + +lemma suminf_cmult_indicator: + fixes f :: "nat \ ereal" + assumes "disjoint_family A" "x \ A i" "\i. 0 \ f i" + shows "(\n. f n * indicator (A n) x) = f i" +proof - + have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)" + using `x \ A i` assms unfolding disjoint_family_on_def indicator_def by auto + then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" + from this[of "Suc i"] show "f i \ y" by auto + qed (insert assms, simp) + ultimately show ?thesis using assms + by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def) +qed + +lemma suminf_indicator: + assumes "disjoint_family A" + shows "(\n. indicator (A n) x :: ereal) = indicator (\i. A i) x" +proof cases + assume *: "x \ (\i. A i)" + then obtain i where "x \ A i" by auto + from suminf_cmult_indicator[OF assms(1), OF `x \ A i`, of "\k. 1"] + show ?thesis using * by simp +qed simp + +text {* + The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to + represent sigma algebras (with an arbitrary emeasure). +*} + +section "Extend binary sets" + +lemma LIMSEQ_binaryset: + assumes f: "f {} = 0" + shows "(\n. \i f A + f B" +proof - + have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" + proof + fix n + show "(\i < Suc (Suc n). f (binaryset A B i)) = f A + f B" + by (induct n) (auto simp add: binaryset_def f) + qed + moreover + have "... ----> f A + f B" by (rule tendsto_const) + ultimately + have "(\n. \i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B" + by metis + hence "(\n. \i< n+2. f (binaryset A B i)) ----> f A + f B" + by simp + thus ?thesis by (rule LIMSEQ_offset [where k=2]) +qed + +lemma binaryset_sums: + assumes f: "f {} = 0" + shows "(\n. f (binaryset A B n)) sums (f A + f B)" + by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) + +lemma suminf_binaryset_eq: + fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" + shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" + by (metis binaryset_sums sums_unique) + +section {* Properties of a premeasure @{term \} *} + +text {* + The definitions for @{const positive} and @{const countably_additive} should be here, by they are + necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}. +*} + +definition additive where + "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" + +definition increasing where + "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" + +lemma positiveD_empty: + "positive M f \ f {} = 0" + by (auto simp add: positive_def) + +lemma additiveD: + "additive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) = f x + f y" + by (auto simp add: additive_def) + +lemma increasingD: + "increasing M f \ x \ y \ x\M \ y\M \ f x \ f y" + by (auto simp add: increasing_def) + +lemma countably_additiveI: + "(\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (\i. f (A i)) = f (\i. A i)) + \ countably_additive M f" + by (simp add: countably_additive_def) + +lemma (in ring_of_sets) disjointed_additive: + assumes f: "positive M f" "additive M f" and A: "range A \ M" "incseq A" + shows "(\i\n. f (disjointed A i)) = f (A n)" +proof (induct n) + case (Suc n) + then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" + by simp + also have "\ = f (A n \ disjointed A (Suc n))" + using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq) + also have "A n \ disjointed A (Suc n) = A (Suc n)" + using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq) + finally show ?case . +qed simp + +lemma (in ring_of_sets) additive_sum: + fixes A:: "'i \ 'a set" + assumes f: "positive M f" and ad: "additive M f" and "finite S" + and A: "A`S \ M" + and disj: "disjoint_family_on A S" + shows "(\i\S. f (A i)) = f (\i\S. A i)" +using `finite S` disj A proof induct + case empty show ?case using f by (simp add: positive_def) +next + case (insert s S) + then have "A s \ (\i\S. A i) = {}" + by (auto simp add: disjoint_family_on_def neq_iff) + moreover + have "A s \ M" using insert by blast + moreover have "(\i\S. A i) \ M" + using insert `finite S` by auto + moreover + ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" + using ad UNION_in_sets A by (auto simp add: additive_def) + with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] + by (auto simp add: additive_def subset_insertI) +qed + +lemma (in ring_of_sets) additive_increasing: + assumes posf: "positive M f" and addf: "additive M f" + shows "increasing M f" +proof (auto simp add: increasing_def) + fix x y + assume xy: "x \ M" "y \ M" "x \ y" + then have "y - x \ M" by auto + then have "0 \ f (y-x)" using posf[unfolded positive_def] by auto + then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono) auto + also have "... = f (x \ (y-x))" using addf + by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) + also have "... = f y" + by (metis Un_Diff_cancel Un_absorb1 xy(3)) + finally show "f x \ f y" by simp +qed + +lemma (in ring_of_sets) countably_additive_additive: + assumes posf: "positive M f" and ca: "countably_additive M f" + shows "additive M f" +proof (auto simp add: additive_def) + fix x y + assume x: "x \ M" and y: "y \ M" and "x \ y = {}" + hence "disjoint_family (binaryset x y)" + by (auto simp add: disjoint_family_on_def binaryset_def) + hence "range (binaryset x y) \ M \ + (\i. binaryset x y i) \ M \ + f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" + using ca + by (simp add: countably_additive_def) + hence "{x,y,{}} \ M \ x \ y \ M \ + f (x \ y) = (\n. f (binaryset x y n))" + by (simp add: range_binaryset_eq UN_binaryset_eq) + thus "f (x \ y) = f x + f y" using posf x y + by (auto simp add: Un suminf_binaryset_eq positive_def) +qed + +lemma (in algebra) increasing_additive_bound: + fixes A:: "nat \ 'a set" and f :: "'a set \ ereal" + assumes f: "positive M f" and ad: "additive M f" + and inc: "increasing M f" + and A: "range A \ M" + and disj: "disjoint_family A" + shows "(\i. f (A i)) \ f \" +proof (safe intro!: suminf_bound) + fix N + note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. f \" using space_closed A + by (intro increasingD[OF inc] finite_UN) auto + finally show "(\i f \" by simp +qed (insert f A, auto simp: positive_def) + +lemma (in ring_of_sets) countably_additiveI_finite: + assumes "finite \" "positive M \" "additive M \" + shows "countably_additive M \" +proof (rule countably_additiveI) + fix F :: "nat \ 'a set" assume F: "range F \ M" "(\i. F i) \ M" and disj: "disjoint_family F" + + have "\i\{i. F i \ {}}. \x. x \ F i" by auto + from bchoice[OF this] obtain f where f: "\i. F i \ {} \ f i \ F i" by auto + + have inj_f: "inj_on f {i. F i \ {}}" + proof (rule inj_onI, simp) + fix i j a b assume *: "f i = f j" "F i \ {}" "F j \ {}" + then have "f i \ F i" "f j \ F j" using f by force+ + with disj * show "i = j" by (auto simp: disjoint_family_on_def) + qed + have "finite (\i. F i)" + by (metis F(2) assms(1) infinite_super sets_into_space) + + have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" + by (auto simp: positiveD_empty[OF `positive M \`]) + moreover have fin_not_empty: "finite {i. F i \ {}}" + proof (rule finite_imageD) + from f have "f`{i. F i \ {}} \ (\i. F i)" by auto + then show "finite (f`{i. F i \ {}})" + by (rule finite_subset) fact + qed fact + ultimately have fin_not_0: "finite {i. \ (F i) \ 0}" + by (rule finite_subset) + + have disj_not_empty: "disjoint_family_on F {i. F i \ {}}" + using disj by (auto simp: disjoint_family_on_def) + + from fin_not_0 have "(\i. \ (F i)) = (\i | \ (F i) \ 0. \ (F i))" + by (rule suminf_eq_setsum) + also have "\ = (\i | F i \ {}. \ (F i))" + using fin_not_empty F_subset by (rule setsum_mono_zero_left) auto + also have "\ = \ (\i\{i. F i \ {}}. F i)" + using `positive M \` `additive M \` fin_not_empty disj_not_empty F by (intro additive_sum) auto + also have "\ = \ (\i. F i)" + by (rule arg_cong[where f=\]) auto + finally show "(\i. \ (F i)) = \ (\i. F i)" . +qed + +section {* Properties of @{const emeasure} *} + +lemma emeasure_positive: "positive (sets M) (emeasure M)" + by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) + +lemma emeasure_empty[simp, intro]: "emeasure M {} = 0" + using emeasure_positive[of M] by (simp add: positive_def) + +lemma emeasure_nonneg[intro!]: "0 \ emeasure M A" + using emeasure_notin_sets[of A M] emeasure_positive[of M] + by (cases "A \ sets M") (auto simp: positive_def) + +lemma emeasure_not_MInf[simp]: "emeasure M A \ - \" + using emeasure_nonneg[of M A] by auto + +lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" + by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) + +lemma suminf_emeasure: + "range A \ sets M \ disjoint_family A \ (\i. emeasure M (A i)) = emeasure M (\i. A i)" + using countable_UN[of A UNIV M] emeasure_countably_additive[of M] + by (simp add: countably_additive_def) + +lemma emeasure_additive: "additive (sets M) (emeasure M)" + by (metis countably_additive_additive emeasure_positive emeasure_countably_additive) + +lemma plus_emeasure: + "a \ sets M \ b \ sets M \ a \ b = {} \ emeasure M a + emeasure M b = emeasure M (a \ b)" + using additiveD[OF emeasure_additive] .. + +lemma setsum_emeasure: + "F`I \ sets M \ disjoint_family_on F I \ finite I \ + (\i\I. emeasure M (F i)) = emeasure M (\i\I. F i)" + by (metis additive_sum emeasure_positive emeasure_additive) + +lemma emeasure_mono: + "a \ b \ b \ sets M \ emeasure M a \ emeasure M b" + by (metis additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets + emeasure_positive increasingD) + +lemma emeasure_space: + "emeasure M A \ emeasure M (space M)" + by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets_into_space top) + +lemma emeasure_compl: + assumes s: "s \ sets M" and fin: "emeasure M s \ \" + shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s" +proof - + from s have "0 \ emeasure M s" by auto + have "emeasure M (space M) = emeasure M (s \ (space M - s))" using s + by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) + also have "... = emeasure M s + emeasure M (space M - s)" + by (rule plus_emeasure[symmetric]) (auto simp add: s) + finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" . + then show ?thesis + using fin `0 \ emeasure M s` + unfolding ereal_eq_minus_iff by (auto simp: ac_simps) +qed + +lemma emeasure_Diff: + assumes finite: "emeasure M B \ \" + and measurable: "A \ sets M" "B \ sets M" "B \ A" + shows "emeasure M (A - B) = emeasure M A - emeasure M B" +proof - + have "0 \ emeasure M B" using assms by auto + have "(A - B) \ B = A" using `B \ A` by auto + then have "emeasure M A = emeasure M ((A - B) \ B)" by simp + also have "\ = emeasure M (A - B) + emeasure M B" + using measurable by (subst plus_emeasure[symmetric]) auto + finally show "emeasure M (A - B) = emeasure M A - emeasure M B" + unfolding ereal_eq_minus_iff + using finite `0 \ emeasure M B` by auto +qed + +lemma emeasure_countable_increasing: + assumes A: "range A \ sets M" + and A0: "A 0 = {}" + and ASuc: "\n. A n \ A (Suc n)" + shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" +proof - + { fix n + have "emeasure M (A n) = (\i (A (Suc m) - A m)" + by (metis ASuc Un_Diff_cancel Un_absorb1) + hence "emeasure M (A (Suc m)) = + emeasure M (A m) + emeasure M (A (Suc m) - A m)" + by (subst plus_emeasure) + (auto simp add: emeasure_additive range_subsetD [OF A]) + with Suc show ?case + by simp + qed } + note Meq = this + have Aeq: "(\i. A (Suc i) - A i) = (\i. A i)" + proof (rule UN_finite2_eq [where k=1], simp) + fix i + show "(\i\{0..i\{0..i. A (Suc i) - A i \ sets M" + by (metis A Diff range_subsetD) + have A2: "(\i. A (Suc i) - A i) \ sets M" + by (blast intro: range_subsetD [OF A]) + have "(SUP n. \ii. emeasure M (A (Suc i) - A i))" + using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric]) + also have "\ = emeasure M (\i. A (Suc i) - A i)" + by (rule suminf_emeasure) + (auto simp add: disjoint_family_Suc ASuc A1 A2) + also have "... = emeasure M (\i. A i)" + by (simp add: Aeq) + finally have "(SUP n. \ii. A i)" . + then show ?thesis by (auto simp add: Meq) +qed + +lemma SUP_emeasure_incseq: + assumes A: "range A \ sets M" and "incseq A" + shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" +proof - + have *: "(SUP n. emeasure M (nat_case {} A (Suc n))) = (SUP n. emeasure M (nat_case {} A n))" + using A by (auto intro!: SUPR_eq exI split: nat.split) + have ueq: "(\i. nat_case {} A i) = (\i. A i)" + by (auto simp add: split: nat.splits) + have meq: "\n. emeasure M (A n) = (emeasure M \ nat_case {} A) (Suc n)" + by simp + have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\i. nat_case {} A i)" + using range_subsetD[OF A] incseq_SucD[OF `incseq A`] + by (force split: nat.splits intro!: emeasure_countable_increasing) + also have "emeasure M (\i. nat_case {} A i) = emeasure M (\i. A i)" + by (simp add: ueq) + finally have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\i. A i)" . + thus ?thesis unfolding meq * comp_def . +qed + +lemma incseq_emeasure: + assumes "range B \ sets M" "incseq B" + shows "incseq (\i. emeasure M (B i))" + using assms by (auto simp: incseq_def intro!: emeasure_mono) + +lemma Lim_emeasure_incseq: + assumes A: "range A \ sets M" "incseq A" + shows "(\i. (emeasure M (A i))) ----> emeasure M (\i. A i)" + using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A] + SUP_emeasure_incseq[OF A] by simp + +lemma decseq_emeasure: + assumes "range B \ sets M" "decseq B" + shows "decseq (\i. emeasure M (B i))" + using assms by (auto simp: decseq_def intro!: emeasure_mono) + +lemma INF_emeasure_decseq: + assumes A: "range A \ sets M" and "decseq A" + and finite: "\i. emeasure M (A i) \ \" + shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" +proof - + have le_MI: "emeasure M (\i. A i) \ emeasure M (A 0)" + using A by (auto intro!: emeasure_mono) + hence *: "emeasure M (\i. A i) \ \" using finite[of 0] by auto + + have A0: "0 \ emeasure M (A 0)" using A by auto + + have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))" + by (simp add: ereal_SUPR_uminus minus_ereal_def) + also have "\ = (SUP n. emeasure M (A 0) - emeasure M (A n))" + unfolding minus_ereal_def using A0 assms + by (subst SUPR_ereal_add) (auto simp add: decseq_emeasure) + also have "\ = (SUP n. emeasure M (A 0 - A n))" + using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto + also have "\ = emeasure M (\i. A 0 - A i)" + proof (rule SUP_emeasure_incseq) + show "range (\n. A 0 - A n) \ sets M" + using A by auto + show "incseq (\n. A 0 - A n)" + using `decseq A` by (auto simp add: incseq_def decseq_def) + qed + also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" + using A finite * by (simp, subst emeasure_Diff) auto + finally show ?thesis + unfolding ereal_minus_eq_minus_iff using finite A0 by auto +qed + +lemma Lim_emeasure_decseq: + assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" + shows "(\i. emeasure M (A i)) ----> emeasure M (\i. A i)" + using LIMSEQ_ereal_INFI[OF decseq_emeasure, OF A] + using INF_emeasure_decseq[OF A fin] by simp + +lemma emeasure_subadditive: + assumes measurable: "A \ sets M" "B \ sets M" + shows "emeasure M (A \ B) \ emeasure M A + emeasure M B" +proof - + from plus_emeasure[of A M "B - A"] + have "emeasure M (A \ B) = emeasure M A + emeasure M (B - A)" + using assms by (simp add: Diff) + also have "\ \ emeasure M A + emeasure M B" + using assms by (auto intro!: add_left_mono emeasure_mono) + finally show ?thesis . +qed + +lemma emeasure_subadditive_finite: + assumes "finite I" "A ` I \ sets M" + shows "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" +using assms proof induct + case (insert i I) + then have "emeasure M (\i\insert i I. A i) = emeasure M (A i \ (\i\I. A i))" + by simp + also have "\ \ emeasure M (A i) + emeasure M (\i\I. A i)" + using insert by (intro emeasure_subadditive finite_UN) auto + also have "\ \ emeasure M (A i) + (\i\I. emeasure M (A i))" + using insert by (intro add_mono) auto + also have "\ = (\i\insert i I. emeasure M (A i))" + using insert by auto + finally show ?case . +qed simp + +lemma emeasure_subadditive_countably: + assumes "range f \ sets M" + shows "emeasure M (\i. f i) \ (\i. emeasure M (f i))" +proof - + have "emeasure M (\i. f i) = emeasure M (\i. disjointed f i)" + unfolding UN_disjointed_eq .. + also have "\ = (\i. emeasure M (disjointed f i))" + using range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] + by (simp add: disjoint_family_disjointed comp_def) + also have "\ \ (\i. emeasure M (f i))" + using range_disjointed_sets[OF assms] assms + by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset) + finally show ?thesis . +qed + +lemma emeasure_insert: + assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" + shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" +proof - + have "{x} \ A = {}" using `x \ A` by auto + from plus_emeasure[OF sets this] show ?thesis by simp +qed + +lemma emeasure_eq_setsum_singleton: + assumes "finite S" "\x. x \ S \ {x} \ sets M" + shows "emeasure M S = (\x\S. emeasure M {x})" + using setsum_emeasure[of "\x. {x}" S M] assms + by (auto simp: disjoint_family_on_def subset_eq) + +lemma setsum_emeasure_cover: + assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" + assumes A: "A \ (\i\S. B i)" + assumes disj: "disjoint_family_on B S" + shows "emeasure M A = (\i\S. emeasure M (A \ (B i)))" +proof - + have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" + proof (rule setsum_emeasure) + show "disjoint_family_on (\i. A \ B i) S" + using `disjoint_family_on B S` + unfolding disjoint_family_on_def by auto + qed (insert assms, auto) + also have "(\i\S. A \ (B i)) = A" + using A by auto + finally show ?thesis by simp +qed + +lemma emeasure_eq_0: + "N \ sets M \ emeasure M N = 0 \ K \ N \ emeasure M K = 0" + by (metis emeasure_mono emeasure_nonneg order_eq_iff) + +lemma emeasure_UN_eq_0: + assumes "\i::nat. emeasure M (N i) = 0" and "range N \ sets M" + shows "emeasure M (\ i. N i) = 0" +proof - + have "0 \ emeasure M (\ i. N i)" using assms by auto + moreover have "emeasure M (\ i. N i) \ 0" + using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp + ultimately show ?thesis by simp +qed + +lemma measure_eqI_finite: + assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" + assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" + shows "M = N" +proof (rule measure_eqI) + fix X assume "X \ sets M" + then have X: "X \ A" by auto + then have "emeasure M X = (\a\X. emeasure M {a})" + using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) + also have "\ = (\a\X. emeasure N {a})" + using X eq by (auto intro!: setsum_cong) + also have "\ = emeasure N X" + using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset) + finally show "emeasure M X = emeasure N X" . +qed simp + +lemma measure_eqI_generator_eq: + fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \ 'a set" + assumes "Int_stable E" "E \ Pow \" + and eq: "\X. X \ E \ emeasure M X = emeasure N X" + and M: "sets M = sigma_sets \ E" + and N: "sets N = sigma_sets \ E" + and A: "range A \ E" "incseq A" "(\i. A i) = \" "\i. emeasure M (A i) \ \" + shows "M = N" +proof - + let ?D = "\F. {D. D \ sigma_sets \ E \ emeasure M (F \ D) = emeasure N (F \ D)}" + interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact + { fix F assume "F \ E" and "emeasure M F \ \" + then have [intro]: "F \ sigma_sets \ E" by auto + have "emeasure N F \ \" using `emeasure M F \ \` `F \ E` eq by simp + interpret D: dynkin_system \ "?D F" + proof (rule dynkin_systemI, simp_all) + fix A assume "A \ sigma_sets \ E \ emeasure M (F \ A) = emeasure N (F \ A)" + then show "A \ \" using S.sets_into_space by auto + next + have "F \ \ = F" using `F \ E` `E \ Pow \` by auto + then show "emeasure M (F \ \) = emeasure N (F \ \)" + using `F \ E` eq by (auto intro: sigma_sets_top) + next + fix A assume *: "A \ sigma_sets \ E \ emeasure M (F \ A) = emeasure N (F \ A)" + then have **: "F \ (\ - A) = F - (F \ A)" + and [intro]: "F \ A \ sigma_sets \ E" + using `F \ E` S.sets_into_space by auto + have "emeasure N (F \ A) \ emeasure N F" by (auto intro!: emeasure_mono simp: M N) + then have "emeasure N (F \ A) \ \" using `emeasure N F \ \` by auto + have "emeasure M (F \ A) \ emeasure M F" by (auto intro!: emeasure_mono simp: M N) + then have "emeasure M (F \ A) \ \" using `emeasure M F \ \` by auto + then have "emeasure M (F \ (\ - A)) = emeasure M F - emeasure M (F \ A)" unfolding ** + using `F \ A \ sigma_sets \ E` by (auto intro!: emeasure_Diff simp: M N) + also have "\ = emeasure N F - emeasure N (F \ A)" using eq `F \ E` * by simp + also have "\ = emeasure N (F \ (\ - A))" unfolding ** + using `F \ A \ sigma_sets \ E` `emeasure N (F \ A) \ \` + by (auto intro!: emeasure_Diff[symmetric] simp: M N) + finally show "\ - A \ sigma_sets \ E \ emeasure M (F \ (\ - A)) = emeasure N (F \ (\ - A))" + using * by auto + next + fix A :: "nat \ 'a set" + assume "disjoint_family A" "range A \ {X \ sigma_sets \ E. emeasure M (F \ X) = emeasure N (F \ X)}" + then have A: "range (\i. F \ A i) \ sigma_sets \ E" "F \ (\x. A x) = (\x. F \ A x)" + "disjoint_family (\i. F \ A i)" "\i. emeasure M (F \ A i) = emeasure N (F \ A i)" "range A \ sigma_sets \ E" + by (auto simp: disjoint_family_on_def subset_eq) + then show "(\x. A x) \ sigma_sets \ E \ emeasure M (F \ (\x. A x)) = emeasure N (F \ (\x. A x))" + by (auto simp: M N suminf_emeasure[symmetric] simp del: UN_simps) + qed + have *: "sigma_sets \ E = ?D F" + using `F \ E` `Int_stable E` + by (intro D.dynkin_lemma) (auto simp add: Int_stable_def eq) + have "\D. D \ sigma_sets \ E \ emeasure M (F \ D) = emeasure N (F \ D)" + by (subst (asm) *) auto } + note * = this + show "M = N" + proof (rule measure_eqI) + show "sets M = sets N" + using M N by simp + fix X assume "X \ sets M" + then have "X \ sigma_sets \ E" + using M by simp + let ?A = "\i. A i \ X" + have "range ?A \ sigma_sets \ E" "incseq ?A" + using A(1,2) `X \ sigma_sets \ E` by (auto simp: incseq_def) + moreover + { fix i have "emeasure M (?A i) = emeasure N (?A i)" + using *[of "A i" X] `X \ sigma_sets \ E` A finite by auto } + ultimately show "emeasure M X = emeasure N X" + using SUP_emeasure_incseq[of ?A M] SUP_emeasure_incseq[of ?A N] A(3) `X \ sigma_sets \ E` + by (auto simp: M N SUP_emeasure_incseq) + qed +qed + +lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" +proof (intro measure_eqI emeasure_measure_of_sigma) + show "sigma_algebra (space M) (sets M)" .. + show "positive (sets M) (emeasure M)" + by (simp add: positive_def emeasure_nonneg) + show "countably_additive (sets M) (emeasure M)" + by (simp add: emeasure_countably_additive) +qed simp_all + +section "@{text \}-null sets" + +definition null_sets :: "'a measure \ 'a set set" where + "null_sets M = {N\sets M. emeasure M N = 0}" + +lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" + by (simp add: null_sets_def) + +lemma null_setsD2[dest]: "A \ null_sets M \ A \ sets M" + unfolding null_sets_def by simp + +lemma null_setsI[intro]: "emeasure M A = 0 \ A \ sets M \ A \ null_sets M" + unfolding null_sets_def by simp + +interpretation null_sets: ring_of_sets "space M" "null_sets M" for M +proof + show "null_sets M \ Pow (space M)" + using sets_into_space by auto + show "{} \ null_sets M" + by auto + fix A B assume sets: "A \ null_sets M" "B \ null_sets M" + then have "A \ sets M" "B \ sets M" + by auto + moreover then have "emeasure M (A \ B) \ emeasure M A + emeasure M B" + "emeasure M (A - B) \ emeasure M A" + by (auto intro!: emeasure_subadditive emeasure_mono) + moreover have "emeasure M B = 0" "emeasure M A = 0" + using sets by auto + ultimately show "A - B \ null_sets M" "A \ B \ null_sets M" + by (auto intro!: antisym) +qed + +lemma UN_from_nat: "(\i. N i) = (\i. N (Countable.from_nat i))" +proof - + have "(\i. N i) = (\i. (N \ Countable.from_nat) i)" + unfolding SUP_def image_compose + unfolding surj_from_nat .. + then show ?thesis by simp +qed + +lemma null_sets_UN[intro]: + assumes "\i::'i::countable. N i \ null_sets M" + shows "(\i. N i) \ null_sets M" +proof (intro conjI CollectI null_setsI) + show "(\i. N i) \ sets M" using assms by auto + have "0 \ emeasure M (\i. N i)" by (rule emeasure_nonneg) + moreover have "emeasure M (\i. N i) \ (\n. emeasure M (N (Countable.from_nat n)))" + unfolding UN_from_nat[of N] + using assms by (intro emeasure_subadditive_countably) auto + ultimately show "emeasure M (\i. N i) = 0" + using assms by (auto simp: null_setsD1) +qed + +lemma null_set_Int1: + assumes "B \ null_sets M" "A \ sets M" shows "A \ B \ null_sets M" +proof (intro CollectI conjI null_setsI) + show "emeasure M (A \ B) = 0" using assms + by (intro emeasure_eq_0[of B _ "A \ B"]) auto +qed (insert assms, auto) + +lemma null_set_Int2: + assumes "B \ null_sets M" "A \ sets M" shows "B \ A \ null_sets M" + using assms by (subst Int_commute) (rule null_set_Int1) + +lemma emeasure_Diff_null_set: + assumes "B \ null_sets M" "A \ sets M" + shows "emeasure M (A - B) = emeasure M A" +proof - + have *: "A - B = (A - (A \ B))" by auto + have "A \ B \ null_sets M" using assms by (rule null_set_Int1) + then show ?thesis + unfolding * using assms + by (subst emeasure_Diff) auto +qed + +lemma null_set_Diff: + assumes "B \ null_sets M" "A \ sets M" shows "B - A \ null_sets M" +proof (intro CollectI conjI null_setsI) + show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto +qed (insert assms, auto) + +lemma emeasure_Un_null_set: + assumes "A \ sets M" "B \ null_sets M" + shows "emeasure M (A \ B) = emeasure M A" +proof - + have *: "A \ B = A \ (B - A)" by auto + have "B - A \ null_sets M" using assms(2,1) by (rule null_set_Diff) + then show ?thesis + unfolding * using assms + by (subst plus_emeasure[symmetric]) auto +qed + +section "Formalize almost everywhere" + +definition ae_filter :: "'a measure \ 'a filter" where + "ae_filter M = Abs_filter (\P. \N\null_sets M. {x \ space M. \ P x} \ N)" + +abbreviation + almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where + "almost_everywhere M P \ eventually P (ae_filter M)" + +syntax + "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) + +translations + "AE x in M. P" == "CONST almost_everywhere M (%x. P)" + +lemma eventually_ae_filter: + fixes M P + defines [simp]: "F \ \P. \N\null_sets M. {x \ space M. \ P x} \ N" + shows "eventually P (ae_filter M) \ F P" + unfolding ae_filter_def F_def[symmetric] +proof (rule eventually_Abs_filter) + show "is_filter F" + proof + fix P Q assume "F P" "F Q" + then obtain N L where N: "N \ null_sets M" "{x \ space M. \ P x} \ N" + and L: "L \ null_sets M" "{x \ space M. \ Q x} \ L" + by auto + then have "L \ N \ null_sets M" "{x \ space M. \ (P x \ Q x)} \ L \ N" by auto + then show "F (\x. P x \ Q x)" by auto + next + fix P Q assume "F P" + then obtain N where N: "N \ null_sets M" "{x \ space M. \ P x} \ N" by auto + moreover assume "\x. P x \ Q x" + ultimately have "N \ null_sets M" "{x \ space M. \ Q x} \ N" by auto + then show "F Q" by auto + qed auto +qed + +lemma AE_I': + "N \ null_sets M \ {x\space M. \ P x} \ N \ (AE x in M. P x)" + unfolding eventually_ae_filter by auto + +lemma AE_iff_null: + assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") + shows "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" +proof + assume "AE x in M. P x" then obtain N where N: "N \ sets M" "?P \ N" "emeasure M N = 0" + unfolding eventually_ae_filter by auto + have "0 \ emeasure M ?P" by auto + moreover have "emeasure M ?P \ emeasure M N" + using assms N(1,2) by (auto intro: emeasure_mono) + ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto + then show "?P \ null_sets M" using assms by auto +next + assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') +qed + +lemma AE_iff_null_sets: + "N \ sets M \ N \ null_sets M \ (AE x in M. x \ N)" + using Int_absorb1[OF sets_into_space, of N M] + by (subst AE_iff_null) (auto simp: Int_def[symmetric]) + +lemma AE_iff_measurable: + "N \ sets M \ {x\space M. \ P x} = N \ (AE x in M. P x) \ emeasure M N = 0" + using AE_iff_null[of _ P] by auto + +lemma AE_E[consumes 1]: + assumes "AE x in M. P x" + obtains N where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" + using assms unfolding eventually_ae_filter by auto + +lemma AE_E2: + assumes "AE x in M. P x" "{x\space M. P x} \ sets M" + shows "emeasure M {x\space M. \ P x} = 0" (is "emeasure M ?P = 0") +proof - + have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto + with AE_iff_null[of M P] assms show ?thesis by auto +qed + +lemma AE_I: + assumes "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" + shows "AE x in M. P x" + using assms unfolding eventually_ae_filter by auto + +lemma AE_mp[elim!]: + assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \ Q x" + shows "AE x in M. Q x" +proof - + from AE_P obtain A where P: "{x\space M. \ P x} \ A" + and A: "A \ sets M" "emeasure M A = 0" + by (auto elim!: AE_E) + + from AE_imp obtain B where imp: "{x\space M. P x \ \ Q x} \ B" + and B: "B \ sets M" "emeasure M B = 0" + by (auto elim!: AE_E) + + show ?thesis + proof (intro AE_I) + have "0 \ emeasure M (A \ B)" using A B by auto + moreover have "emeasure M (A \ B) \ 0" + using emeasure_subadditive[of A M B] A B by auto + ultimately show "A \ B \ sets M" "emeasure M (A \ B) = 0" using A B by auto + show "{x\space M. \ Q x} \ A \ B" + using P imp by auto + qed +qed + +(* depricated replace by laws about eventually *) +lemma + shows AE_iffI: "AE x in M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" + and AE_disjI1: "AE x in M. P x \ AE x in M. P x \ Q x" + and AE_disjI2: "AE x in M. Q x \ AE x in M. P x \ Q x" + and AE_conjI: "AE x in M. P x \ AE x in M. Q x \ AE x in M. P x \ Q x" + and AE_conj_iff[simp]: "(AE x in M. P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" + by auto + +lemma AE_impI: + "(P \ AE x in M. Q x) \ AE x in M. P \ Q x" + by (cases P) auto + +lemma AE_measure: + assumes AE: "AE x in M. P x" and sets: "{x\space M. P x} \ sets M" (is "?P \ sets M") + shows "emeasure M {x\space M. P x} = emeasure M (space M)" +proof - + from AE_E[OF AE] guess N . note N = this + with sets have "emeasure M (space M) \ emeasure M (?P \ N)" + by (intro emeasure_mono) auto + also have "\ \ emeasure M ?P + emeasure M N" + using sets N by (intro emeasure_subadditive) auto + also have "\ = emeasure M ?P" using N by simp + finally show "emeasure M ?P = emeasure M (space M)" + using emeasure_space[of M "?P"] by auto +qed + +lemma AE_space: "AE x in M. x \ space M" + by (rule AE_I[where N="{}"]) auto + +lemma AE_I2[simp, intro]: + "(\x. x \ space M \ P x) \ AE x in M. P x" + using AE_space by force + +lemma AE_Ball_mp: + "\x\space M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" + by auto + +lemma AE_cong[cong]: + "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" + by auto + +lemma AE_all_countable: + "(AE x in M. \i. P i x) \ (\i::'i::countable. AE x in M. P i x)" +proof + assume "\i. AE x in M. P i x" + from this[unfolded eventually_ae_filter Bex_def, THEN choice] + obtain N where N: "\i. N i \ null_sets M" "\i. {x\space M. \ P i x} \ N i" by auto + have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto + also have "\ \ (\i. N i)" using N by auto + finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . + moreover from N have "(\i. N i) \ null_sets M" + by (intro null_sets_UN) auto + ultimately show "AE x in M. \i. P i x" + unfolding eventually_ae_filter by auto +qed auto + +lemma AE_finite_all: + assumes f: "finite S" shows "(AE x in M. \i\S. P i x) \ (\i\S. AE x in M. P i x)" + using f by induct auto + +lemma AE_finite_allI: + assumes "finite S" + shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" + using AE_finite_all[OF `finite S`] by auto + +lemma emeasure_mono_AE: + assumes imp: "AE x in M. x \ A \ x \ B" + and B: "B \ sets M" + shows "emeasure M A \ emeasure M B" +proof cases + assume A: "A \ sets M" + from imp obtain N where N: "{x\space M. \ (x \ A \ x \ B)} \ N" "N \ null_sets M" + by (auto simp: eventually_ae_filter) + have "emeasure M A = emeasure M (A - N)" + using N A by (subst emeasure_Diff_null_set) auto + also have "emeasure M (A - N) \ emeasure M (B - N)" + using N A B sets_into_space by (auto intro!: emeasure_mono) + also have "emeasure M (B - N) = emeasure M B" + using N B by (subst emeasure_Diff_null_set) auto + finally show ?thesis . +qed (simp add: emeasure_nonneg emeasure_notin_sets) + +lemma emeasure_eq_AE: + assumes iff: "AE x in M. x \ A \ x \ B" + assumes A: "A \ sets M" and B: "B \ sets M" + shows "emeasure M A = emeasure M B" + using assms by (safe intro!: antisym emeasure_mono_AE) auto + +section {* @{text \}-finite Measures *} + +locale sigma_finite_measure = + fixes M :: "'a measure" + assumes sigma_finite: "\A::nat \ 'a set. + range A \ sets M \ (\i. A i) = space M \ (\i. emeasure M (A i) \ \)" + +lemma (in sigma_finite_measure) sigma_finite_disjoint: + obtains A :: "nat \ 'a set" + where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" +proof atomize_elim + case goal1 + obtain A :: "nat \ 'a set" where + range: "range A \ sets M" and + space: "(\i. A i) = space M" and + measure: "\i. emeasure M (A i) \ \" + using sigma_finite by auto + note range' = range_disjointed_sets[OF range] range + { fix i + have "emeasure M (disjointed A i) \ emeasure M (A i)" + using range' disjointed_subset[of A i] by (auto intro!: emeasure_mono) + then have "emeasure M (disjointed A i) \ \" + using measure[of i] by auto } + with disjoint_family_disjointed UN_disjointed_eq[of A] space range' + show ?case by (auto intro!: exI[of _ "disjointed A"]) +qed + +lemma (in sigma_finite_measure) sigma_finite_incseq: + obtains A :: "nat \ 'a set" + where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" +proof atomize_elim + case goal1 + obtain F :: "nat \ 'a set" where + F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" + using sigma_finite by auto + then show ?case + proof (intro exI[of _ "\n. \i\n. F i"] conjI allI) + from F have "\x. x \ space M \ \i. x \ F i" by auto + then show "(\n. \ i\n. F i) = space M" + using F by fastforce + next + fix n + have "emeasure M (\ i\n. F i) \ (\i\n. emeasure M (F i))" using F + by (auto intro!: emeasure_subadditive_finite) + also have "\ < \" + using F by (auto simp: setsum_Pinfty) + finally show "emeasure M (\ i\n. F i) \ \" by simp + qed (force simp: incseq_def)+ +qed + +section {* Measure space induced by distribution of @{const measurable}-functions *} + +definition distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where + "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" + +lemma + shows sets_distr[simp]: "sets (distr M N f) = sets N" + and space_distr[simp]: "space (distr M N f) = space N" + by (auto simp: distr_def) + +lemma + shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" + and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" + by (auto simp: measurable_def) + +lemma emeasure_distr: + fixes f :: "'a \ 'b" + assumes f: "f \ measurable M N" and A: "A \ sets N" + shows "emeasure (distr M N f) A = emeasure M (f -` A \ space M)" (is "_ = ?\ A") + unfolding distr_def +proof (rule emeasure_measure_of_sigma) + show "positive (sets N) ?\" + by (auto simp: positive_def) + + show "countably_additive (sets N) ?\" + proof (intro countably_additiveI) + fix A :: "nat \ 'b set" assume "range A \ sets N" "disjoint_family A" + then have A: "\i. A i \ sets N" "(\i. A i) \ sets N" by auto + then have *: "range (\i. f -` (A i) \ space M) \ sets M" + using f by (auto simp: measurable_def) + moreover have "(\i. f -` A i \ space M) \ sets M" + using * by blast + moreover have **: "disjoint_family (\i. f -` A i \ space M)" + using `disjoint_family A` by (auto simp: disjoint_family_on_def) + ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" + using suminf_emeasure[OF _ **] A f + by (auto simp: comp_def vimage_UN) + qed + show "sigma_algebra (space N) (sets N)" .. +qed fact + +lemma AE_distrD: + assumes f: "f \ measurable M M'" + and AE: "AE x in distr M M' f. P x" + shows "AE x in M. P (f x)" +proof - + from AE[THEN AE_E] guess N . + with f show ?thesis + unfolding eventually_ae_filter + by (intro bexI[of _ "f -` N \ space M"]) + (auto simp: emeasure_distr measurable_def) +qed + +lemma null_sets_distr_iff: + "f \ measurable M N \ A \ null_sets (distr M N f) \ f -` A \ space M \ null_sets M \ A \ sets N" + by (auto simp add: null_sets_def emeasure_distr measurable_sets) + +lemma distr_distr: + assumes f: "g \ measurable N L" and g: "f \ measurable M N" + shows "distr (distr M N f) L g = distr M L (g \ f)" (is "?L = ?R") + using measurable_comp[OF g f] f g + by (auto simp add: emeasure_distr measurable_sets measurable_space + intro!: arg_cong[where f="emeasure M"] measure_eqI) + +section {* Real measure values *} + +lemma measure_nonneg: "0 \ measure M A" + using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos) + +lemma measure_empty[simp]: "measure M {} = 0" + unfolding measure_def by simp + +lemma emeasure_eq_ereal_measure: + "emeasure M A \ \ \ emeasure M A = ereal (measure M A)" + using emeasure_nonneg[of M A] + by (cases "emeasure M A") (auto simp: measure_def) + +lemma measure_Union: + assumes finite: "emeasure M A \ \" "emeasure M B \ \" + and measurable: "A \ sets M" "B \ sets M" "A \ B = {}" + shows "measure M (A \ B) = measure M A + measure M B" + unfolding measure_def + using plus_emeasure[OF measurable, symmetric] finite + by (simp add: emeasure_eq_ereal_measure) + +lemma measure_finite_Union: + assumes measurable: "A`S \ sets M" "disjoint_family_on A S" "finite S" + assumes finite: "\i. i \ S \ emeasure M (A i) \ \" + shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" + unfolding measure_def + using setsum_emeasure[OF measurable, symmetric] finite + by (simp add: emeasure_eq_ereal_measure) + +lemma measure_Diff: + assumes finite: "emeasure M A \ \" + and measurable: "A \ sets M" "B \ sets M" "B \ A" + shows "measure M (A - B) = measure M A - measure M B" +proof - + have "emeasure M (A - B) \ emeasure M A" "emeasure M B \ emeasure M A" + using measurable by (auto intro!: emeasure_mono) + hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" + using measurable finite by (rule_tac measure_Union) auto + thus ?thesis using `B \ A` by (auto simp: Un_absorb2) +qed + +lemma measure_UNION: + assumes measurable: "range A \ sets M" "disjoint_family A" + assumes finite: "emeasure M (\i. A i) \ \" + shows "(\i. measure M (A i)) sums (measure M (\i. A i))" +proof - + from summable_sums[OF summable_ereal_pos, of "\i. emeasure M (A i)"] + suminf_emeasure[OF measurable] emeasure_nonneg[of M] + have "(\i. emeasure M (A i)) sums (emeasure M (\i. A i))" by simp + moreover + { fix i + have "emeasure M (A i) \ emeasure M (\i. A i)" + using measurable by (auto intro!: emeasure_mono) + then have "emeasure M (A i) = ereal ((measure M (A i)))" + using finite by (intro emeasure_eq_ereal_measure) auto } + ultimately show ?thesis using finite + unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure) +qed + +lemma measure_subadditive: + assumes measurable: "A \ sets M" "B \ sets M" + and fin: "emeasure M A \ \" "emeasure M B \ \" + shows "(measure M (A \ B)) \ (measure M A) + (measure M B)" +proof - + have "emeasure M (A \ B) \ \" + using emeasure_subadditive[OF measurable] fin by auto + then show "(measure M (A \ B)) \ (measure M A) + (measure M B)" + using emeasure_subadditive[OF measurable] fin + by (auto simp: emeasure_eq_ereal_measure) +qed + +lemma measure_subadditive_finite: + assumes A: "finite I" "A`I \ sets M" and fin: "\i. i \ I \ emeasure M (A i) \ \" + shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" +proof - + { have "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" + using emeasure_subadditive_finite[OF A] . + also have "\ < \" + using fin by (simp add: setsum_Pinfty) + finally have "emeasure M (\i\I. A i) \ \" by simp } + then show ?thesis + using emeasure_subadditive_finite[OF A] fin + unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg) +qed + +lemma measure_subadditive_countably: + assumes A: "range A \ sets M" and fin: "(\i. emeasure M (A i)) \ \" + shows "measure M (\i. A i) \ (\i. measure M (A i))" +proof - + from emeasure_nonneg fin have "\i. emeasure M (A i) \ \" by (rule suminf_PInfty) + moreover + { have "emeasure M (\i. A i) \ (\i. emeasure M (A i))" + using emeasure_subadditive_countably[OF A] . + also have "\ < \" + using fin by simp + finally have "emeasure M (\i. A i) \ \" by simp } + ultimately show ?thesis + using emeasure_subadditive_countably[OF A] fin + unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg) +qed + +lemma measure_eq_setsum_singleton: + assumes S: "finite S" "\x. x \ S \ {x} \ sets M" + and fin: "\x. x \ S \ emeasure M {x} \ \" + shows "(measure M S) = (\x\S. (measure M {x}))" + unfolding measure_def + using emeasure_eq_setsum_singleton[OF S] fin + by simp (simp add: emeasure_eq_ereal_measure) + +lemma Lim_measure_incseq: + assumes A: "range A \ sets M" "incseq A" and fin: "emeasure M (\i. A i) \ \" + shows "(\i. (measure M (A i))) ----> (measure M (\i. A i))" +proof - + have "ereal ((measure M (\i. A i))) = emeasure M (\i. A i)" + using fin by (auto simp: emeasure_eq_ereal_measure) + then show ?thesis + using Lim_emeasure_incseq[OF A] + unfolding measure_def + by (intro lim_real_of_ereal) simp +qed + +lemma Lim_measure_decseq: + assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" + shows "(\n. measure M (A n)) ----> measure M (\i. A i)" +proof - + have "emeasure M (\i. A i) \ emeasure M (A 0)" + using A by (auto intro!: emeasure_mono) + also have "\ < \" + using fin[of 0] by auto + finally have "ereal ((measure M (\i. A i))) = emeasure M (\i. A i)" + by (auto simp: emeasure_eq_ereal_measure) + then show ?thesis + unfolding measure_def + using Lim_emeasure_decseq[OF A fin] + by (intro lim_real_of_ereal) simp +qed + +section {* Measure spaces with @{term "emeasure M (space M) < \"} *} + +locale finite_measure = sigma_finite_measure M for M + + assumes finite_emeasure_space: "emeasure M (space M) \ \" + +lemma finite_measureI[Pure.intro!]: + assumes *: "emeasure M (space M) \ \" + shows "finite_measure M" +proof + show "\A. range A \ sets M \ (\i. A i) = space M \ (\i. emeasure M (A i) \ \)" + using * by (auto intro!: exI[of _ "\_. space M"]) +qed fact + +lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \ \" + using finite_emeasure_space emeasure_space[of M A] by auto + +lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)" + unfolding measure_def by (simp add: emeasure_eq_ereal_measure) + +lemma (in finite_measure) emeasure_real: "\r. 0 \ r \ emeasure M A = ereal r" + using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto + +lemma (in finite_measure) bounded_measure: "measure M A \ measure M (space M)" + using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) + +lemma (in finite_measure) finite_measure_Diff: + assumes sets: "A \ sets M" "B \ sets M" and "B \ A" + shows "measure M (A - B) = measure M A - measure M B" + using measure_Diff[OF _ assms] by simp + +lemma (in finite_measure) finite_measure_Union: + assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" + shows "measure M (A \ B) = measure M A + measure M B" + using measure_Union[OF _ _ assms] by simp + +lemma (in finite_measure) finite_measure_finite_Union: + assumes measurable: "A`S \ sets M" "disjoint_family_on A S" "finite S" + shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" + using measure_finite_Union[OF assms] by simp + +lemma (in finite_measure) finite_measure_UNION: + assumes A: "range A \ sets M" "disjoint_family A" + shows "(\i. measure M (A i)) sums (measure M (\i. A i))" + using measure_UNION[OF A] by simp + +lemma (in finite_measure) finite_measure_mono: + assumes "A \ B" "B \ sets M" shows "measure M A \ measure M B" + using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) + +lemma (in finite_measure) finite_measure_subadditive: + assumes m: "A \ sets M" "B \ sets M" + shows "measure M (A \ B) \ measure M A + measure M B" + using measure_subadditive[OF m] by simp + +lemma (in finite_measure) finite_measure_subadditive_finite: + assumes "finite I" "A`I \ sets M" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" + using measure_subadditive_finite[OF assms] by simp + +lemma (in finite_measure) finite_measure_subadditive_countably: + assumes A: "range A \ sets M" and sum: "summable (\i. measure M (A i))" + shows "measure M (\i. A i) \ (\i. measure M (A i))" +proof - + from `summable (\i. measure M (A i))` + have "(\i. ereal (measure M (A i))) sums ereal (\i. measure M (A i))" + by (simp add: sums_ereal) (rule summable_sums) + from sums_unique[OF this, symmetric] + measure_subadditive_countably[OF A] + show ?thesis by (simp add: emeasure_eq_measure) +qed + +lemma (in finite_measure) finite_measure_eq_setsum_singleton: + assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" + shows "measure M S = (\x\S. measure M {x})" + using measure_eq_setsum_singleton[OF assms] by simp + +lemma (in finite_measure) finite_Lim_measure_incseq: + assumes A: "range A \ sets M" "incseq A" + shows "(\i. measure M (A i)) ----> measure M (\i. A i)" + using Lim_measure_incseq[OF A] by simp + +lemma (in finite_measure) finite_Lim_measure_decseq: + assumes A: "range A \ sets M" "decseq A" + shows "(\n. measure M (A n)) ----> measure M (\i. A i)" + using Lim_measure_decseq[OF A] by simp + +lemma (in finite_measure) finite_measure_compl: + assumes S: "S \ sets M" + shows "measure M (space M - S) = measure M (space M) - measure M S" + using measure_Diff[OF _ top S sets_into_space] S by simp + +lemma (in finite_measure) finite_measure_mono_AE: + assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" + shows "measure M A \ measure M B" + using assms emeasure_mono_AE[OF imp B] + by (simp add: emeasure_eq_measure) + +lemma (in finite_measure) finite_measure_eq_AE: + assumes iff: "AE x in M. x \ A \ x \ B" + assumes A: "A \ sets M" and B: "B \ sets M" + shows "measure M A = measure M B" + using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) + +section {* Counting space *} + +definition count_space :: "'a set \ 'a measure" where + "count_space \ = measure_of \ (Pow \) (\A. if finite A then ereal (card A) else \)" + +lemma + shows space_count_space[simp]: "space (count_space \) = \" + and sets_count_space[simp]: "sets (count_space \) = Pow \" + using sigma_sets_into_sp[of "Pow \" \] + by (auto simp: count_space_def) + +lemma measurable_count_space_eq1[simp]: + "f \ measurable (count_space A) M \ f \ A \ space M" + unfolding measurable_def by simp + +lemma measurable_count_space_eq2[simp]: + assumes "finite A" + shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" +proof - + { fix X assume "X \ A" "f \ space M \ A" + with `finite A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "finite X" + by (auto dest: finite_subset) + moreover assume "\a\A. f -` {a} \ space M \ sets M" + ultimately have "f -` X \ space M \ sets M" + using `X \ A` by (auto intro!: finite_UN simp del: UN_simps) } + then show ?thesis + unfolding measurable_def by auto +qed + +lemma emeasure_count_space: + assumes "X \ A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \)" + (is "_ = ?M X") + unfolding count_space_def +proof (rule emeasure_measure_of_sigma) + show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) + + show "positive (Pow A) ?M" + by (auto simp: positive_def) + + show "countably_additive (Pow A) ?M" + proof (unfold countably_additive_def, safe) + fix F :: "nat \ 'a set" assume disj: "disjoint_family F" + show "(\i. ?M (F i)) = ?M (\i. F i)" + proof cases + assume "\i. finite (F i)" + then have finite_F: "\i. finite (F i)" by auto + have "\i\{i. F i \ {}}. \x. x \ F i" by auto + from bchoice[OF this] obtain f where f: "\i. F i \ {} \ f i \ F i" by auto + + have inj_f: "inj_on f {i. F i \ {}}" + proof (rule inj_onI, simp) + fix i j a b assume *: "f i = f j" "F i \ {}" "F j \ {}" + then have "f i \ F i" "f j \ F j" using f by force+ + with disj * show "i = j" by (auto simp: disjoint_family_on_def) + qed + have fin_eq: "finite (\i. F i) \ finite {i. F i \ {}}" + proof + assume "finite (\i. F i)" + show "finite {i. F i \ {}}" + proof (rule finite_imageD) + from f have "f`{i. F i \ {}} \ (\i. F i)" by auto + then show "finite (f`{i. F i \ {}})" + by (rule finite_subset) fact + qed fact + next + assume "finite {i. F i \ {}}" + with finite_F have "finite (\i\{i. F i \ {}}. F i)" + by auto + also have "(\i\{i. F i \ {}}. F i) = (\i. F i)" + by auto + finally show "finite (\i. F i)" . + qed + + show ?thesis + proof cases + assume *: "finite (\i. F i)" + with finite_F have "finite {i. ?M (F i) \ 0} " + by (simp add: fin_eq) + then have "(\i. ?M (F i)) = (\i | ?M (F i) \ 0. ?M (F i))" + by (rule suminf_eq_setsum) + also have "\ = ereal (\i | F i \ {}. card (F i))" + using finite_F by simp + also have "\ = ereal (card (\i \ {i. F i \ {}}. F i))" + using * finite_F disj + by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def fin_eq) + also have "\ = ?M (\i. F i)" + using * by (auto intro!: arg_cong[where f=card]) + finally show ?thesis . + next + assume inf: "infinite (\i. F i)" + { fix i + have "\N. i \ (\i (\i {}} - {..< N})" + using inf by (auto simp: fin_eq) + then have "{i. F i \ {}} - {..< N} \ {}" + by (metis finite.emptyI) + then obtain i where i: "F i \ {}" "N \ i" + by (auto simp: not_less[symmetric]) + + note N + also have "(\i (\i < (\i {}` by (simp add: card_gt_0_iff) + finally have "j < (\i (\i. finite (F i))" + then obtain j where j: "infinite (F j)" by auto + then have "infinite (\i. F i)" + using finite_subset[of "F j" "\i. F i"] by auto + moreover have "\i. 0 \ ?M (F i)" by auto + ultimately show ?thesis + using suminf_PInfty[of "\i. ?M (F i)" j] j by auto + qed + qed + show "X \ Pow A" using `X \ A` by simp +qed + +lemma emeasure_count_space_finite[simp]: + "X \ A \ finite X \ emeasure (count_space A) X = ereal (card X)" + using emeasure_count_space[of X A] by simp + +lemma emeasure_count_space_infinite[simp]: + "X \ A \ infinite X \ emeasure (count_space A) X = \" + using emeasure_count_space[of X A] by simp + +lemma emeasure_count_space_eq_0: + "emeasure (count_space A) X = 0 \ (X \ A \ X = {})" +proof cases + assume X: "X \ A" + then show ?thesis + proof (intro iffI impI) + assume "emeasure (count_space A) X = 0" + with X show "X = {}" + by (subst (asm) emeasure_count_space) (auto split: split_if_asm) + qed simp +qed (simp add: emeasure_notin_sets) + +lemma null_sets_count_space: "null_sets (count_space A) = { {} }" + unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) + +lemma AE_count_space: "(AE x in count_space A. P x) \ (\x\A. P x)" + unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) + +lemma sigma_finite_measure_count_space: + fixes A :: "'a::countable set" + shows "sigma_finite_measure (count_space A)" +proof + show "\F::nat \ 'a set. range F \ sets (count_space A) \ (\i. F i) = space (count_space A) \ + (\i. emeasure (count_space A) (F i) \ \)" + using surj_from_nat by (intro exI[of _ "\i. {from_nat i} \ A"]) (auto simp del: surj_from_nat) +qed + +lemma finite_measure_count_space: + assumes [simp]: "finite A" + shows "finite_measure (count_space A)" + by rule simp + +lemma sigma_finite_measure_count_space_finite: + assumes A: "finite A" shows "sigma_finite_measure (count_space A)" +proof - + interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) + show "sigma_finite_measure (count_space A)" .. +qed + +end + diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Probability.thy Mon Apr 23 12:14:35 2012 +0200 @@ -4,7 +4,7 @@ Probability_Measure Infinite_Product_Measure Independent_Family - Conditional_Probability Information begin + end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Probability_Measure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -6,110 +6,219 @@ header {*Probability measure*} theory Probability_Measure -imports Lebesgue_Measure + imports Lebesgue_Measure Radon_Nikodym begin +lemma funset_eq_UN_fun_upd_I: + assumes *: "\f. f \ F (insert a A) \ f(a := d) \ F A" + and **: "\f. f \ F (insert a A) \ f a \ G (f(a:=d))" + and ***: "\f x. \ f \ F A ; x \ G f \ \ f(a:=x) \ F (insert a A)" + shows "F (insert a A) = (\f\F A. fun_upd f a ` (G f))" +proof safe + fix f assume f: "f \ F (insert a A)" + show "f \ (\f\F A. fun_upd f a ` G f)" + proof (rule UN_I[of "f(a := d)"]) + show "f(a := d) \ F A" using *[OF f] . + show "f \ fun_upd (f(a:=d)) a ` G (f(a:=d))" + proof (rule image_eqI[of _ _ "f a"]) + show "f a \ G (f(a := d))" using **[OF f] . + qed simp + qed +next + fix f x assume "f \ F A" "x \ G f" + from ***[OF this] show "f(a := x) \ F (insert a A)" . +qed + +lemma extensional_funcset_insert_eq[simp]: + assumes "a \ A" + shows "extensional (insert a A) \ (insert a A \ B) = (\f \ extensional A \ (A \ B). (\b. f(a := b)) ` B)" + apply (rule funset_eq_UN_fun_upd_I) + using assms + by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def) + +lemma finite_extensional_funcset[simp, intro]: + assumes "finite A" "finite B" + shows "finite (extensional A \ (A \ B))" + using assms by induct auto + +lemma finite_PiE[simp, intro]: + assumes fin: "finite A" "\i. i \ A \ finite (B i)" + shows "finite (Pi\<^isub>E A B)" +proof - + have *: "(Pi\<^isub>E A B) \ extensional A \ (A \ (\i\A. B i))" by auto + show ?thesis + using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto +qed + + +lemma countably_additiveI[case_names countably]: + assumes "\A. \ range A \ M ; disjoint_family A ; (\i. A i) \ M\ \ (\n. \ (A n)) = \ (\i. A i)" + shows "countably_additive M \" + using assms unfolding countably_additive_def by auto + +lemma convex_le_Inf_differential: + fixes f :: "real \ real" + assumes "convex_on I f" + assumes "x \ interior I" "y \ I" + shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" + (is "_ \ _ + Inf (?F x) * (y - x)") +proof - + show ?thesis + proof (cases rule: linorder_cases) + assume "x < y" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "min (x + e / 2) ((x + y) / 2)" + ultimately have "x < t" "t < y" "t \ ball x e" + by (auto simp: dist_real_def field_simps split: split_min) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . + moreover def K \ "x - e / 2" + with `0 < e` have "K \ ball x e" "K < x" by (auto simp: dist_real_def) + ultimately have "K \ I" "K < x" "x \ I" + using interior_subset[of I] `x \ interior I` by auto + + have "Inf (?F x) \ (f x - f y) / (x - y)" + proof (rule Inf_lower2) + show "(f x - f t) / (x - t) \ ?F x" + using `t \ I` `x < t` by auto + show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + using `convex_on I f` `x \ I` `y \ I` `x < t` `t < y` by (rule convex_on_diff) + next + fix y assume "y \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `K \ I` _ `K < x` _]] + show "(f K - f x) / (K - x) \ y" by auto + qed + then show ?thesis + using `x < y` by (simp add: field_simps) + next + assume "y < x" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "x + e / 2" + ultimately have "x < t" "t \ ball x e" + by (auto simp: dist_real_def field_simps) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "(f x - f y) / (x - y) \ Inf (?F x)" + proof (rule Inf_greatest) + have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" + using `y < x` by (auto simp: field_simps) + also + fix z assume "z \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `y \ I` _ `y < x`]] + have "(f y - f x) / (y - x) \ z" by auto + finally show "(f x - f y) / (x - y) \ z" . + next + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) + with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto + then show "?F x \ {}" by blast + qed + then show ?thesis + using `y < x` by (simp add: field_simps) + qed simp +qed + +lemma distr_id[simp]: "distr N N (\x. x) = N" + by (rule measure_eqI) (auto simp: emeasure_distr) + locale prob_space = finite_measure + - assumes measure_space_1: "measure M (space M) = 1" + assumes emeasure_space_1: "emeasure M (space M) = 1" lemma prob_spaceI[Pure.intro!]: - assumes "measure_space M" - assumes *: "measure M (space M) = 1" + assumes *: "emeasure M (space M) = 1" shows "prob_space M" proof - interpret finite_measure M proof - show "measure_space M" by fact - show "measure M (space M) \ \" using * by simp + show "emeasure M (space M) \ \" using * by simp qed show "prob_space M" by default fact qed abbreviation (in prob_space) "events \ sets M" -abbreviation (in prob_space) "prob \ \'" -abbreviation (in prob_space) "random_variable M' X \ sigma_algebra M' \ X \ measurable M M'" +abbreviation (in prob_space) "prob \ measure M" +abbreviation (in prob_space) "random_variable M' X \ X \ measurable M M'" abbreviation (in prob_space) "expectation \ integral\<^isup>L M" -definition (in prob_space) - "distribution X A = \' (X -` A \ space M)" - -abbreviation (in prob_space) - "joint_distribution X Y \ distribution (\x. (X x, Y x))" - -lemma (in prob_space) prob_space_cong: - assumes "\A. A \ sets M \ measure N A = \ A" "space N = space M" "sets N = sets M" - shows "prob_space N" -proof - show "measure_space N" by (intro measure_space_cong assms) - show "measure N (space N) = 1" - using measure_space_1 assms by simp +lemma (in prob_space) prob_space_distr: + assumes f: "f \ measurable M M'" shows "prob_space (distr M M' f)" +proof (rule prob_spaceI) + have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) + with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1" + by (auto simp: emeasure_distr emeasure_space_1) qed -lemma (in prob_space) distribution_cong: - assumes "\x. x \ space M \ X x = Y x" - shows "distribution X = distribution Y" - unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\'"]) - -lemma (in prob_space) joint_distribution_cong: - assumes "\x. x \ space M \ X x = X' x" - assumes "\x. x \ space M \ Y x = Y' x" - shows "joint_distribution X Y = joint_distribution X' Y'" - unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\'"]) - -lemma (in prob_space) distribution_id[simp]: - "N \ events \ distribution (\x. x) N = prob N" - by (auto simp: distribution_def intro!: arg_cong[where f=prob]) - lemma (in prob_space) prob_space: "prob (space M) = 1" - using measure_space_1 unfolding \'_def by (simp add: one_ereal_def) + using emeasure_space_1 unfolding measure_def by (simp add: one_ereal_def) lemma (in prob_space) prob_le_1[simp, intro]: "prob A \ 1" using bounded_measure[of A] by (simp add: prob_space) -lemma (in prob_space) distribution_positive[simp, intro]: - "0 \ distribution X A" unfolding distribution_def by auto - -lemma (in prob_space) not_zero_less_distribution[simp]: - "(\ 0 < distribution X A) \ distribution X A = 0" - using distribution_positive[of X A] by arith - -lemma (in prob_space) joint_distribution_remove[simp]: - "joint_distribution X X {(x, x)} = distribution X {x}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) +lemma (in prob_space) not_empty: "space M \ {}" + using prob_space by auto -lemma (in prob_space) distribution_unit[simp]: "distribution (\x. ()) {()} = 1" - unfolding distribution_def using prob_space by auto - -lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\x. (X x, ())) {(a, ())} = distribution X {a}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) - -lemma (in prob_space) not_empty: "space M \ {}" - using prob_space empty_measure' by auto - -lemma (in prob_space) measure_le_1: "X \ sets M \ \ X \ 1" - unfolding measure_space_1[symmetric] - using sets_into_space - by (intro measure_mono) auto +lemma (in prob_space) measure_le_1: "emeasure M X \ 1" + using emeasure_space[of M X] by (simp add: emeasure_space_1) lemma (in prob_space) AE_I_eq_1: - assumes "\ {x\space M. P x} = 1" "{x\space M. P x} \ sets M" - shows "AE x. P x" + assumes "emeasure M {x\space M. P x} = 1" "{x\space M. P x} \ sets M" + shows "AE x in M. P x" proof (rule AE_I) - show "\ (space M - {x \ space M. P x}) = 0" - using assms measure_space_1 by (simp add: measure_compl) + show "emeasure M (space M - {x \ space M. P x}) = 0" + using assms emeasure_space_1 by (simp add: emeasure_compl) qed (insert assms, auto) -lemma (in prob_space) distribution_1: - "distribution X A \ 1" - unfolding distribution_def by simp - lemma (in prob_space) prob_compl: assumes A: "A \ events" shows "prob (space M - A) = 1 - prob A" using finite_measure_compl[OF A] by (simp add: prob_space) +lemma (in prob_space) AE_in_set_eq_1: + assumes "A \ events" shows "(AE x in M. x \ A) \ prob A = 1" +proof + assume ae: "AE x in M. x \ A" + have "{x \ space M. x \ A} = A" "{x \ space M. x \ A} = space M - A" + using `A \ events`[THEN sets_into_space] by auto + with AE_E2[OF ae] `A \ events` have "1 - emeasure M A = 0" + by (simp add: emeasure_compl emeasure_space_1) + then show "prob A = 1" + using `A \ events` by (simp add: emeasure_eq_measure one_ereal_def) +next + assume prob: "prob A = 1" + show "AE x in M. x \ A" + proof (rule AE_I) + show "{x \ space M. x \ A} \ space M - A" by auto + show "emeasure M (space M - A) = 0" + using `A \ events` prob + by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def) + show "space M - A \ events" + using `A \ events` by auto + qed +qed + +lemma (in prob_space) AE_False: "(AE x in M. False) \ False" +proof + assume "AE x in M. False" + then have "AE x in M. x \ {}" by simp + then show False + by (subst (asm) AE_in_set_eq_1) auto +qed simp + +lemma (in prob_space) AE_prob_1: + assumes "prob A = 1" shows "AE x in M. x \ A" +proof - + from `prob A = 1` have "A \ events" + by (metis measure_notin_sets zero_neq_one) + with AE_in_set_eq_1 assms show ?thesis by simp +qed + lemma (in prob_space) prob_space_increasing: "increasing M prob" by (auto intro!: finite_measure_mono simp: increasing_def) @@ -164,9 +273,8 @@ shows "prob (\ i :: nat. c i) = 0" proof (rule antisym) show "prob (\ i :: nat. c i) \ 0" - using finite_measure_countably_subadditive[OF assms(1)] - by (simp add: assms(2) suminf_zero summable_zero) -qed simp + using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) +qed (simp add: measure_nonneg) lemma (in prob_space) prob_equiprobable_finite_unions: assumes "s \ events" @@ -178,7 +286,7 @@ from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ prob {x} = prob {SOME y. y \ s}" by blast have "prob s = (\ x \ s. prob {x})" - using finite_measure_finite_singleton[OF s_finite] by simp + using finite_measure_eq_setsum_singleton[OF s_finite] by simp also have "\ = (\ x \ s. prob {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * prob {(SOME x. x \ s)}" using setsum_constant assms by (simp add: real_eq_of_nat) @@ -199,96 +307,20 @@ also have "\ = (\ x \ s. prob (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact - show "\i. i \ s \ e \ f i \ events" by fact + show "(\i. e \ f i)`s \ events" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed -lemma (in prob_space) prob_space_vimage: - assumes S: "sigma_algebra S" - assumes T: "T \ measure_preserving M S" - shows "prob_space S" -proof - interpret S: measure_space S - using S and T by (rule measure_space_vimage) - show "measure_space S" .. - - from T[THEN measure_preservingD2] - have "T -` space S \ space M = space M" - by (auto simp: measurable_def) - with T[THEN measure_preservingD, of "space S", symmetric] - show "measure S (space S) = 1" - using measure_space_1 by simp -qed - -lemma prob_space_unique_Int_stable: - fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \ 'a set" - assumes E: "Int_stable E" "space E \ sets E" - and M: "prob_space M" "space M = space E" "sets M = sets (sigma E)" - and N: "prob_space N" "space N = space E" "sets N = sets (sigma E)" - and eq: "\X. X \ sets E \ finite_measure.\' M X = finite_measure.\' N X" - assumes "X \ sets (sigma E)" - shows "finite_measure.\' M X = finite_measure.\' N X" -proof - - interpret M!: prob_space M by fact - interpret N!: prob_space N by fact - have "measure M X = measure N X" - proof (rule measure_unique_Int_stable[OF `Int_stable E`]) - show "range (\i. space M) \ sets E" "incseq (\i. space M)" "(\i. space M) = space E" - using E M N by auto - show "\i. M.\ (space M) \ \" - using M.measure_space_1 by simp - show "measure_space \space = space E, sets = sets (sigma E), measure_space.measure = M.\\" - using E M N by (auto intro!: M.measure_space_cong) - show "measure_space \space = space E, sets = sets (sigma E), measure_space.measure = N.\\" - using E M N by (auto intro!: N.measure_space_cong) - { fix X assume "X \ sets E" - then have "X \ sets (sigma E)" - by (auto simp: sets_sigma sigma_sets.Basic) - with eq[OF `X \ sets E`] M N show "M.\ X = N.\ X" - by (simp add: M.finite_measure_eq N.finite_measure_eq) } - qed fact - with `X \ sets (sigma E)` M N show ?thesis - by (simp add: M.finite_measure_eq N.finite_measure_eq) -qed - -lemma (in prob_space) distribution_prob_space: - assumes X: "random_variable S X" - shows "prob_space (S\measure := ereal \ distribution X\)" (is "prob_space ?S") -proof (rule prob_space_vimage) - show "X \ measure_preserving M ?S" - using X - unfolding measure_preserving_def distribution_def [abs_def] - by (auto simp: finite_measure_eq measurable_sets) - show "sigma_algebra ?S" using X by simp -qed - -lemma (in prob_space) AE_distribution: - assumes X: "random_variable MX X" and "AE x in MX\measure := ereal \ distribution X\. Q x" - shows "AE x. Q (X x)" -proof - - interpret X: prob_space "MX\measure := ereal \ 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 - from X[unfolded measurable_def] N show "AE x. Q (X x)" - by (intro AE_I'[where N="X -` N \ space M"]) - (auto simp: finite_measure_eq distribution_def measurable_sets) -qed - -lemma (in prob_space) distribution_eq_integral: - "random_variable S X \ A \ sets S \ distribution X A = expectation (indicator (X -` A \ space M))" - using finite_measure_eq[of "X -` A \ space M"] - by (auto simp: measurable_sets distribution_def) - lemma (in prob_space) expectation_less: assumes [simp]: "integrable M X" assumes gt: "\x\space M. X x < b" shows "expectation X < b" proof - have "expectation X < expectation (\x. b)" - using gt measure_space_1 + using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed @@ -299,80 +331,11 @@ shows "a < expectation X" proof - have "expectation (\x. a) < expectation X" - using gt measure_space_1 + using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed -lemma convex_le_Inf_differential: - fixes f :: "real \ real" - assumes "convex_on I f" - assumes "x \ interior I" "y \ I" - shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" - (is "_ \ _ + Inf (?F x) * (y - x)") -proof - - show ?thesis - proof (cases rule: linorder_cases) - assume "x < y" - moreover - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - moreover def t \ "min (x + e / 2) ((x + y) / 2)" - ultimately have "x < t" "t < y" "t \ ball x e" - by (auto simp: mem_ball dist_real_def field_simps split: split_min) - with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto - - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . - moreover def K \ "x - e / 2" - with `0 < e` have "K \ ball x e" "K < x" by (auto simp: mem_ball dist_real_def) - ultimately have "K \ I" "K < x" "x \ I" - using interior_subset[of I] `x \ interior I` by auto - - have "Inf (?F x) \ (f x - f y) / (x - y)" - proof (rule Inf_lower2) - show "(f x - f t) / (x - t) \ ?F x" - using `t \ I` `x < t` by auto - show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" - using `convex_on I f` `x \ I` `y \ I` `x < t` `t < y` by (rule convex_on_diff) - next - fix y assume "y \ ?F x" - with order_trans[OF convex_on_diff[OF `convex_on I f` `K \ I` _ `K < x` _]] - show "(f K - f x) / (K - x) \ y" by auto - qed - then show ?thesis - using `x < y` by (simp add: field_simps) - next - assume "y < x" - moreover - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - moreover def t \ "x + e / 2" - ultimately have "x < t" "t \ ball x e" - by (auto simp: mem_ball dist_real_def field_simps) - with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto - - have "(f x - f y) / (x - y) \ Inf (?F x)" - proof (rule Inf_greatest) - have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" - using `y < x` by (auto simp: field_simps) - also - fix z assume "z \ ?F x" - with order_trans[OF convex_on_diff[OF `convex_on I f` `y \ I` _ `y < x`]] - have "(f y - f x) / (y - x) \ z" by auto - finally show "(f x - f y) / (x - y) \ z" . - next - have "open (interior I)" by auto - from openE[OF this `x \ interior I`] guess e . note e = this - then have "x + e / 2 \ ball x e" by (auto simp: mem_ball dist_real_def) - with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto - then show "?F x \ {}" by blast - qed - then show ?thesis - using `y < x` by (simp add: field_simps) - qed simp -qed - lemma (in prob_space) jensens_inequality: fixes a b :: real assumes X: "integrable M X" "X ` space M \ I" @@ -410,8 +373,7 @@ fix k assume "k \ (\x. q x + ?F x * (expectation X - x)) ` I" then guess x .. note x = this have "q x + ?F x * (expectation X - x) = expectation (\w. q x + ?F x * (X w - x))" - using prob_space - by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X) + using prob_space by (simp add: X) also have "\ \ expectation (\w. q (X w))" using `x \ I` `open I` X(2) by (intro integral_mono integral_add integral_cmult integral_diff @@ -422,31 +384,6 @@ finally show "q (expectation X) \ expectation (\x. q (X x))" . qed -lemma (in prob_space) distribution_eq_translated_integral: - assumes "random_variable S X" "A \ sets S" - shows "distribution X A = integral\<^isup>P (S\measure := ereal \ distribution X\) (indicator A)" -proof - - interpret S: prob_space "S\measure := ereal \ distribution X\" - using assms(1) by (rule distribution_prob_space) - show ?thesis - using S.positive_integral_indicator(1)[of A] assms by simp -qed - -lemma (in prob_space) finite_expectation1: - assumes f: "finite (X`space M)" and rv: "random_variable borel X" - shows "expectation X = (\r \ X ` space M. r * prob (X -` {r} \ space M))" (is "_ = ?r") -proof (subst integral_on_finite) - show "X \ borel_measurable M" "finite (X`space M)" using assms by auto - show "(\ r \ X ` space M. r * real (\ (X -` {r} \ space M))) = ?r" - "\x. \ (X -` {x} \ space M) \ \" - using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto -qed - -lemma (in prob_space) finite_expectation: - assumes "finite (X`space M)" "random_variable borel X" - shows "expectation X = (\ r \ X ` (space M). r * distribution X {r})" - using assms unfolding distribution_def using finite_expectation1 by auto - lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0: assumes "{x} \ events" assumes "prob {x} = 1" @@ -455,119 +392,25 @@ shows "prob {y} = 0" using prob_one_inter[of "{y}" "{x}"] assms by auto -lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0" - unfolding distribution_def by simp - -lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1" -proof - - have "X -` X ` space M \ space M = space M" by auto - thus ?thesis unfolding distribution_def by (simp add: prob_space) -qed - -lemma (in prob_space) distribution_one: - assumes "random_variable M' X" and "A \ sets M'" - shows "distribution X A \ 1" -proof - - have "distribution X A \ \' (space M)" unfolding distribution_def - using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono) - thus ?thesis by (simp add: prob_space) -qed - -lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0: - assumes X: "random_variable \space = X ` (space M), sets = Pow (X ` (space M))\ X" - (is "random_variable ?S X") - assumes "distribution X {x} = 1" - assumes "y \ x" - shows "distribution X {y} = 0" -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 finite_measure_inter_full_set[OF single single, of x y] assms(2) - show ?thesis by (auto simp: distribution_def prob_space) -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: - assumes X: "random_variable MX X" and Y: "random_variable MY Y" - and A: "A \ sets MX" and B: "B \ sets MY" - shows "joint_distribution X Y (A \ B) \ distribution X A" - unfolding distribution_def -proof (intro finite_measure_mono) - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force - show "X -` A \ space M \ events" - using X A unfolding measurable_def by simp - have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = - (X -` A \ space M) \ (Y -` B \ space M)" by auto -qed - -lemma (in prob_space) joint_distribution_commute: - "joint_distribution X Y x = joint_distribution Y X ((\(x,y). (y,x))`x)" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) + "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY + \ emeasure (distr M (MX \\<^isub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MX X) A" + by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) lemma (in prob_space) joint_distribution_Times_le_snd: - assumes X: "random_variable MX X" and Y: "random_variable MY Y" - and A: "A \ sets MX" and B: "B \ sets MY" - shows "joint_distribution X Y (A \ B) \ distribution Y B" - using assms - by (subst joint_distribution_commute) - (simp add: swap_product joint_distribution_Times_le_fst) - -lemma (in prob_space) random_variable_pairI: - assumes "random_variable MX X" - assumes "random_variable MY Y" - 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 (MX \\<^isub>M MY)" by default - have sa: "sigma_algebra M" by default - 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 - -lemma (in prob_space) joint_distribution_commute_singleton: - "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) - -lemma (in prob_space) joint_distribution_assoc_singleton: - "joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = - joint_distribution (\x. (X x, Y x)) Z {((x, y), z)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\']) + "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY + \ emeasure (distr M (MX \\<^isub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MY Y) B" + by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2 -sublocale pair_prob_space \ P: prob_space P +sublocale pair_prob_space \ P: prob_space "M1 \\<^isub>M M2" proof - show "measure_space P" .. - show "measure P (space P) = 1" - by (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure) + show "emeasure (M1 \\<^isub>M M2) (space (M1 \\<^isub>M M2)) = 1" + by (simp add: emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure) qed -lemma countably_additiveI[case_names countably]: - assumes "\A. \ range A \ sets M ; disjoint_family A ; (\i. A i) \ sets M\ \ - (\n. \ (A n)) = \ (\i. A i)" - shows "countably_additive M \" - using assms unfolding countably_additive_def by auto - -lemma (in prob_space) joint_distribution_prob_space: - assumes "random_variable MX X" "random_variable MY Y" - shows "prob_space ((MX \\<^isub>M MY) \ measure := ereal \ joint_distribution X Y\)" - using random_variable_pairI[OF assms] by (rule distribution_prob_space) - -locale product_prob_space = product_sigma_finite M for M :: "'i \ ('a, 'b) measure_space_scheme" + +locale product_prob_space = product_sigma_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes prob_space: "\i. prob_space (M i)" @@ -578,648 +421,401 @@ sublocale finite_product_prob_space \ prob_space "\\<^isub>M i\I. M i" proof - show "measure_space P" .. - show "measure P (space P) = 1" - by (simp add: measure_times M.measure_space_1 setprod_1) + show "emeasure (\\<^isub>M i\I. M i) (space (\\<^isub>M i\I. M i)) = 1" + by (simp add: measure_times M.emeasure_space_1 setprod_1 space_PiM) qed lemma (in finite_product_prob_space) prob_times: assumes X: "\i. i \ I \ X i \ sets (M i)" shows "prob (\\<^isub>E i\I. X i) = (\i\I. M.prob i (X i))" proof - - have "ereal (\' (\\<^isub>E i\I. X i)) = \ (\\<^isub>E i\I. X i)" - using X by (intro finite_measure_eq[symmetric] in_P) auto - also have "\ = (\i\I. M.\ i (X i))" + have "ereal (measure (\\<^isub>M i\I. M i) (\\<^isub>E i\I. X i)) = emeasure (\\<^isub>M i\I. M i) (\\<^isub>E i\I. X i)" + using X by (simp add: emeasure_eq_measure) + also have "\ = (\i\I. emeasure (M i) (X i))" using measure_times X by simp - also have "\ = ereal (\i\I. M.\' i (X i))" - using X by (simp add: M.finite_measure_eq setprod_ereal) - finally show ?thesis by simp -qed - -lemma (in prob_space) random_variable_restrict: - assumes I: "finite I" - assumes X: "\i. i \ I \ random_variable (N i) (X i)" - shows "random_variable (Pi\<^isub>M I N) (\x. \i\I. X i x)" -proof - { fix i assume "i \ I" - with X interpret N: sigma_algebra "N i" by simp - have "sets (N i) \ Pow (space (N i))" by (rule N.space_closed) } - note N_closed = this - then show "sigma_algebra (Pi\<^isub>M I N)" - by (simp add: product_algebra_def) - (intro sigma_algebra_sigma product_algebra_generator_sets_into_space) - show "(\x. \i\I. X i x) \ measurable M (Pi\<^isub>M I N)" - using X by (intro measurable_restrict[OF I N_closed]) auto -qed - -section "Probability spaces on finite sets" - -locale finite_prob_space = prob_space + finite_measure_space - -abbreviation (in prob_space) "finite_random_variable M' X \ finite_sigma_algebra M' \ X \ measurable M M'" - -lemma (in prob_space) finite_random_variableD: - assumes "finite_random_variable M' X" shows "random_variable M' X" -proof - - interpret M': finite_sigma_algebra M' using assms by simp - show "random_variable M' X" using assms by simp default -qed - -lemma (in prob_space) distribution_finite_prob_space: - assumes "finite_random_variable MX X" - shows "finite_prob_space (MX\measure := ereal \ distribution X\)" -proof - - interpret X: prob_space "MX\measure := ereal \ distribution X\" - using assms[THEN finite_random_variableD] by (rule distribution_prob_space) - interpret MX: finite_sigma_algebra MX - using assms by auto - show ?thesis by default (simp_all add: MX.finite_space) -qed - -lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]: - 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 ?X by (rule sigma_algebra_Pow) - show "finite_sigma_algebra ?X" - by default auto - 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 - then have "X -` (\a\A. {a}) \ space M \ events" - unfolding vimage_UN UN_extend_simps - apply (rule finite_UN) - using A assms unfolding simple_function_def by auto - then show "X -` A \ space M \ events" by simp - qed -qed - -lemma (in prob_space) simple_function_imp_random_variable[simp, intro]: - 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 M X \ (\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def prob_space[symmetric] - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def simple_function_def - intro!: arg_cong[where f=prob]) - -lemma (in prob_space) finite_random_variable_pairI: - assumes "finite_random_variable MX X" - assumes "finite_random_variable MY Y" - 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 (MX \\<^isub>M MY)" .. - have sa: "sigma_algebra M" by default - 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 - -lemma (in prob_space) finite_random_variable_imp_sets: - "finite_random_variable MX X \ x \ space MX \ {x} \ sets MX" - unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp - -lemma (in prob_space) finite_random_variable_measurable: - assumes X: "finite_random_variable MX X" shows "X -` A \ space M \ events" -proof - - interpret X: finite_sigma_algebra MX using X by simp - from X have vimage: "\A. A \ space MX \ X -` A \ space M \ events" and - "X \ space M \ space MX" - by (auto simp: measurable_def) - then have *: "X -` A \ space M = X -` (A \ space MX) \ space M" - by auto - show "X -` A \ space M \ events" - unfolding * by (intro vimage) auto -qed - -lemma (in prob_space) joint_distribution_finite_Times_le_fst: - assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" - shows "joint_distribution X Y (A \ B) \ distribution X A" - unfolding distribution_def -proof (intro finite_measure_mono) - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force - show "X -` A \ space M \ events" - using finite_random_variable_measurable[OF X] . - have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = - (X -` A \ space M) \ (Y -` B \ space M)" by auto -qed - -lemma (in prob_space) joint_distribution_finite_Times_le_snd: - assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" - shows "joint_distribution X Y (A \ B) \ distribution Y B" - using assms - by (subst joint_distribution_commute) - (simp add: swap_product joint_distribution_finite_Times_le_fst) - -lemma (in prob_space) finite_distribution_order: - fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme" - assumes "finite_random_variable MX X" "finite_random_variable MY Y" - shows "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution X {x}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" - and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" - and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"] - using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"] - by (auto intro: antisym) - -lemma (in prob_space) setsum_joint_distribution: - assumes X: "finite_random_variable MX X" - assumes Y: "random_variable MY Y" "B \ sets MY" - shows "(\a\space MX. joint_distribution X Y ({a} \ B)) = distribution Y B" - unfolding distribution_def -proof (subst finite_measure_finite_Union[symmetric]) - interpret MX: finite_sigma_algebra MX using X by auto - show "finite (space MX)" using MX.finite_space . - let ?d = "\i. (\x. (X x, Y x)) -` ({i} \ B) \ space M" - { fix i assume "i \ space MX" - moreover have "?d i = (X -` {i} \ space M) \ (Y -` B \ space M)" by auto - ultimately show "?d i \ events" - using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y - using MX.sets_eq_Pow by auto } - show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def) - show "\' (\i\space MX. ?d i) = \' (Y -` B \ space M)" - using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\']) -qed - -lemma (in prob_space) setsum_joint_distribution_singleton: - assumes X: "finite_random_variable MX X" - assumes Y: "finite_random_variable MY Y" "b \ space MY" - shows "(\a\space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}" - using setsum_joint_distribution[OF X - finite_random_variableD[OF Y(1)] - finite_random_variable_imp_sets[OF Y]] by simp - -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 - -locale pair_finite_prob_space = pair_prob_space M1 M2 + pair_finite_space M1 M2 + M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 - -sublocale pair_finite_prob_space \ finite_prob_space P by default - -lemma funset_eq_UN_fun_upd_I: - assumes *: "\f. f \ F (insert a A) \ f(a := d) \ F A" - and **: "\f. f \ F (insert a A) \ f a \ G (f(a:=d))" - and ***: "\f x. \ f \ F A ; x \ G f \ \ f(a:=x) \ F (insert a A)" - shows "F (insert a A) = (\f\F A. fun_upd f a ` (G f))" -proof safe - fix f assume f: "f \ F (insert a A)" - show "f \ (\f\F A. fun_upd f a ` G f)" - proof (rule UN_I[of "f(a := d)"]) - show "f(a := d) \ F A" using *[OF f] . - show "f \ fun_upd (f(a:=d)) a ` G (f(a:=d))" - proof (rule image_eqI[of _ _ "f a"]) - show "f a \ G (f(a := d))" using **[OF f] . - qed simp - qed -next - fix f x assume "f \ F A" "x \ G f" - from ***[OF this] show "f(a := x) \ F (insert a A)" . -qed - -lemma extensional_funcset_insert_eq[simp]: - assumes "a \ A" - shows "extensional (insert a A) \ (insert a A \ B) = (\f \ extensional A \ (A \ B). (\b. f(a := b)) ` B)" - apply (rule funset_eq_UN_fun_upd_I) - using assms - by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def) - -lemma finite_extensional_funcset[simp, intro]: - assumes "finite A" "finite B" - shows "finite (extensional A \ (A \ B))" - using assms by induct (auto simp: extensional_funcset_insert_eq) - -lemma finite_PiE[simp, intro]: - assumes fin: "finite A" "\i. i \ A \ finite (B i)" - shows "finite (Pi\<^isub>E A B)" -proof - - have *: "(Pi\<^isub>E A B) \ extensional A \ (A \ (\i\A. B i))" by auto - show ?thesis - using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto -qed - -locale finite_product_finite_prob_space = finite_product_prob_space M I for M I + - assumes finite_space: "\i. finite_prob_space (M i)" - -sublocale finite_product_finite_prob_space \ M!: finite_prob_space "M i" using finite_space . - -lemma (in finite_product_finite_prob_space) singleton_eq_product: - assumes x: "x \ space P" shows "{x} = (\\<^isub>E i\I. {x i})" -proof (safe intro!: ext[of _ x]) - fix y i assume *: "y \ (\ i\I. {x i})" "y \ extensional I" - with x show "y i = x i" - by (cases "i \ I") (auto simp: extensional_def) -qed (insert x, auto) - -sublocale finite_product_finite_prob_space \ finite_prob_space "Pi\<^isub>M I M" -proof - show "finite (space P)" - using finite_index M.finite_space by auto - - { fix x assume "x \ space P" - with this[THEN singleton_eq_product] - have "{x} \ sets P" - by (auto intro!: in_P) } - note x_in_P = this - - have "Pow (space P) \ sets P" - proof - fix X assume "X \ Pow (space P)" - moreover then have "finite X" - using `finite (space P)` by (blast intro: finite_subset) - ultimately have "(\x\X. {x}) \ sets P" - by (intro finite_UN x_in_P) auto - then show "X \ sets P" by simp - qed - with space_closed show [simp]: "sets P = Pow (space P)" .. -qed - -lemma (in finite_product_finite_prob_space) measure_finite_times: - "(\i. i \ I \ X i \ space (M i)) \ \ (\\<^isub>E i\I. X i) = (\i\I. M.\ i (X i))" - by (rule measure_times) simp - -lemma (in finite_product_finite_prob_space) measure_singleton_times: - assumes x: "x \ space P" shows "\ {x} = (\i\I. M.\ i {x i})" - unfolding singleton_eq_product[OF x] using x - by (intro measure_finite_times) auto - -lemma (in finite_product_finite_prob_space) prob_finite_times: - assumes X: "\i. i \ I \ X i \ space (M i)" - shows "prob (\\<^isub>E i\I. X i) = (\i\I. M.prob i (X i))" -proof - - have "ereal (\' (\\<^isub>E i\I. X i)) = \ (\\<^isub>E i\I. X i)" - using X by (intro finite_measure_eq[symmetric] in_P) auto - also have "\ = (\i\I. M.\ i (X i))" - using measure_finite_times X by simp - also have "\ = ereal (\i\I. M.\' i (X i))" - using X by (simp add: M.finite_measure_eq setprod_ereal) + also have "\ = ereal (\i\I. measure (M i) (X i))" + using X by (simp add: M.emeasure_eq_measure setprod_ereal) finally show ?thesis by simp qed -lemma (in finite_product_finite_prob_space) prob_singleton_times: - assumes x: "x \ space P" - shows "prob {x} = (\i\I. M.prob i {x i})" - unfolding singleton_eq_product[OF x] using x - by (intro prob_finite_times) auto - -lemma (in finite_product_finite_prob_space) prob_finite_product: - "A \ space P \ prob A = (\x\A. \i\I. M.prob i {x i})" - by (auto simp add: finite_measure_singleton prob_singleton_times - simp del: space_product_algebra - intro!: setsum_cong prob_singleton_times) +section {* Distributions *} -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 ((MX \\<^isub>M MY)\ measure := ereal \ 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 \ measure M (space M) = 1" - unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def - by auto - -lemma (in finite_prob_space) sum_over_space_eq_1: "(\x\space M. \ {x}) = 1" - using measure_space_1 sum_over_space by simp +definition "distributed M N X f \ distr M N X = density N f \ + f \ borel_measurable N \ (AE x in N. 0 \ f x) \ X \ measurable M N" -lemma (in finite_prob_space) joint_distribution_restriction_fst: - "joint_distribution X Y A \ distribution X (fst ` A)" - unfolding distribution_def -proof (safe intro!: finite_measure_mono) - fix x assume "x \ space M" and *: "(X x, Y x) \ A" - show "x \ X -` fst ` A" - by (auto intro!: image_eqI[OF _ *]) -qed (simp_all add: sets_eq_Pow) - -lemma (in finite_prob_space) joint_distribution_restriction_snd: - "joint_distribution X Y A \ distribution Y (snd ` A)" - unfolding distribution_def -proof (safe intro!: finite_measure_mono) - fix x assume "x \ space M" and *: "(X x, Y x) \ A" - show "x \ Y -` snd ` A" - by (auto intro!: image_eqI[OF _ *]) -qed (simp_all add: sets_eq_Pow) - -lemma (in finite_prob_space) distribution_order: - shows "0 \ distribution X x'" - and "(distribution X x' \ 0) \ (0 < distribution X x')" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution X {x}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" - and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" - and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using - joint_distribution_restriction_fst[of X Y "{(x, y)}"] - joint_distribution_restriction_snd[of X Y "{(x, y)}"] - by (auto intro: antisym) - -lemma (in finite_prob_space) distribution_mono: - assumes "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "distribution X x \ distribution Y y" - unfolding distribution_def - using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono) +lemma + shows distributed_distr_eq_density: "distributed M N X f \ distr M N X = density N f" + and distributed_measurable: "distributed M N X f \ X \ measurable M N" + and distributed_borel_measurable: "distributed M N X f \ f \ borel_measurable N" + and distributed_AE: "distributed M N X f \ (AE x in N. 0 \ f x)" + by (simp_all add: distributed_def) -lemma (in finite_prob_space) distribution_mono_gt_0: - assumes gt_0: "0 < distribution X x" - assumes *: "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "0 < distribution Y y" - by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) - -lemma (in finite_prob_space) sum_over_space_distrib: - "(\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def prob_space[symmetric] using finite_space - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow - intro!: arg_cong[where f=\']) - -lemma (in finite_prob_space) finite_sum_over_space_eq_1: - "(\x\space M. prob {x}) = 1" - using prob_space finite_space - by (subst (asm) finite_measure_finite_singleton) auto - -lemma (in prob_space) distribution_remove_const: - shows "joint_distribution X (\x. ()) {(x, ())} = distribution X {x}" - and "joint_distribution (\x. ()) X {((), x)} = distribution X {x}" - and "joint_distribution X (\x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" - and "joint_distribution X (\x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" - and "distribution (\x. ()) {()} = 1" - by (auto intro!: arg_cong[where f=\'] simp: distribution_def prob_space[symmetric]) +lemma + shows distributed_real_measurable: "distributed M N X (\x. ereal (f x)) \ f \ borel_measurable N" + and distributed_real_AE: "distributed M N X (\x. ereal (f x)) \ (AE x in N. 0 \ f x)" + by (simp_all add: distributed_def borel_measurable_ereal_iff) -lemma (in finite_prob_space) setsum_distribution_gen: - assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" - and "inj_on f (X`space M)" - shows "(\x \ X`space M. distribution Y {f x}) = distribution Z {c}" - unfolding distribution_def assms - using finite_space assms - by (subst finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_prob_space) setsum_distribution_cut: - "(\x \ X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" - "(\y \ Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" - "(\x \ X`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" - "(\y \ Y`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" - "(\z \ Z`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" - by (auto intro!: inj_onI setsum_distribution_gen) - -lemma (in finite_prob_space) uniform_prob: - assumes "x \ space M" - assumes "\ x y. \x \ space M ; y \ space M\ \ prob {x} = prob {y}" - shows "prob {x} = 1 / card (space M)" +lemma distributed_count_space: + assumes X: "distributed M (count_space A) X P" and a: "a \ A" and A: "finite A" + shows "P a = emeasure M (X -` {a} \ space M)" proof - - have prob_x: "\ y. y \ space M \ prob {y} = prob {x}" - using assms(2)[OF _ `x \ space M`] by blast - have "1 = prob (space M)" - using prob_space by auto - also have "\ = (\ x \ space M. prob {x})" - using finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] - sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] - finite_space unfolding disjoint_family_on_def prob_space[symmetric] - by (auto simp add:setsum_restrict_set) - also have "\ = (\ y \ space M. prob {x})" - using prob_x by auto - also have "\ = real_of_nat (card (space M)) * prob {x}" by simp - finally have one: "1 = real (card (space M)) * prob {x}" - using real_eq_of_nat by auto - hence two: "real (card (space M)) \ 0" by fastforce - from one have three: "prob {x} \ 0" by fastforce - thus ?thesis using one two three divide_cancel_right - by (auto simp:field_simps) + have "emeasure M (X -` {a} \ space M) = emeasure (distr M (count_space A) X) {a}" + using X a A by (simp add: distributed_measurable emeasure_distr) + also have "\ = emeasure (density (count_space A) P) {a}" + using X by (simp add: distributed_distr_eq_density) + also have "\ = (\\<^isup>+x. P a * indicator {a} x \count_space A)" + using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: positive_integral_cong) + also have "\ = P a" + using X a by (subst positive_integral_cmult_indicator) (auto simp: distributed_def one_ereal_def[symmetric] AE_count_space) + finally show ?thesis .. qed -lemma (in prob_space) prob_space_subalgebra: - assumes "sigma_algebra N" "sets N \ sets M" "space N = space M" - and "\A. A \ sets N \ measure N A = \ A" - shows "prob_space N" -proof - interpret N: measure_space N - by (rule measure_space_subalgebra[OF assms]) - show "measure_space N" .. - show "measure N (space N) = 1" - using assms(4)[OF N.top] by (simp add: assms measure_space_1) +lemma distributed_cong_density: + "(AE x in N. f x = g x) \ g \ borel_measurable N \ f \ borel_measurable N \ + distributed M N X f \ distributed M N X g" + by (auto simp: distributed_def intro!: density_cong) + +lemma subdensity: + assumes T: "T \ measurable P Q" + assumes f: "distributed M P X f" + assumes g: "distributed M Q Y g" + assumes Y: "Y = T \ X" + shows "AE x in P. g (T x) = 0 \ f x = 0" +proof - + have "{x\space Q. g x = 0} \ null_sets (distr M Q (T \ X))" + using g Y by (auto simp: null_sets_density_iff distributed_def) + also have "distr M Q (T \ X) = distr (distr M P X) Q T" + using T f[THEN distributed_measurable] by (rule distr_distr[symmetric]) + finally have "T -` {x\space Q. g x = 0} \ space P \ null_sets (distr M P X)" + using T by (subst (asm) null_sets_distr_iff) auto + also have "T -` {x\space Q. g x = 0} \ space P = {x\space P. g (T x) = 0}" + using T by (auto dest: measurable_space) + finally show ?thesis + using f g by (auto simp add: null_sets_density_iff distributed_def) qed -lemma (in prob_space) prob_space_of_restricted_space: - assumes "\ A \ 0" "A \ sets M" - shows "prob_space (restricted_space A \measure := \S. \ S / \ A\)" - (is "prob_space ?P") +lemma subdensity_real: + fixes g :: "'a \ real" and f :: "'b \ real" + assumes T: "T \ measurable P Q" + assumes f: "distributed M P X f" + assumes g: "distributed M Q Y g" + assumes Y: "Y = T \ X" + shows "AE x in P. g (T x) = 0 \ f x = 0" + using subdensity[OF T, of M X "\x. ereal (f x)" Y "\x. ereal (g x)"] assms by auto + +lemma distributed_integral: + "distributed M N X f \ g \ borel_measurable N \ (\x. f x * g x \N) = (\x. g (X x) \M)" + by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable + distributed_distr_eq_density[symmetric] integral_density[symmetric] integral_distr) + +lemma distributed_transform_integral: + assumes Px: "distributed M N X Px" + assumes "distributed M P Y Py" + assumes Y: "Y = T \ X" and T: "T \ measurable N P" and f: "f \ borel_measurable P" + shows "(\x. Py x * f x \P) = (\x. Px x * f (T x) \N)" proof - - interpret A: measure_space "restricted_space A" - using `A \ sets M` by (rule restricted_measure_space) - interpret A': sigma_algebra ?P - by (rule A.sigma_algebra_cong) auto - show "prob_space ?P" - proof - show "measure_space ?P" - proof - show "positive ?P (measure ?P)" - proof (simp add: positive_def, safe) - fix B assume "B \ events" - with real_measure[of "A \ B"] real_measure[OF `A \ events`] `A \ sets M` - show "0 \ \ (A \ B) / \ A" by (auto simp: Int) - qed - show "countably_additive ?P (measure ?P)" - proof (simp add: countably_additive_def, safe) - fix B and F :: "nat \ 'a set" - assume F: "range F \ op \ A ` events" "disjoint_family F" - { fix i - from F have "F i \ op \ A ` events" by auto - with `A \ events` have "F i \ events" by auto } - moreover then have "range F \ events" by auto - moreover have "\S. \ S / \ A = inverse (\ A) * \ S" - by (simp add: mult_commute divide_ereal_def) - moreover have "0 \ inverse (\ A)" - using real_measure[OF `A \ events`] by auto - ultimately show "(\i. \ (F i) / \ A) = \ (\i. F i) / \ A" - using measure_countably_additive[of F] F - by (auto simp: suminf_cmult_ereal) - qed - qed - show "measure ?P (space ?P) = 1" - using real_measure[OF `A \ events`] `\ A \ 0` by auto - qed -qed - -lemma finite_prob_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" - and 1: "measure M (space M) = 1" and "\x. x \ space M \ 0 \ measure M {x}" - and add: "\A B. A \ space M \ measure M A = (\x\A. measure M {x})" - shows "finite_prob_space M" -proof - - interpret finite_measure_space M - proof - show "measure M (space M) \ \" using 1 by simp - qed fact+ - show ?thesis by default fact -qed - -lemma (in finite_prob_space) distribution_eq_setsum: - "distribution X A = (\x\A \ X ` space M. distribution X {x})" -proof - - have *: "X -` A \ space M = (\x\A \ X ` space M. X -` {x} \ space M)" - by auto - then show "distribution X A = (\x\A \ X ` space M. distribution X {x})" - using finite_space unfolding distribution_def * - by (intro finite_measure_finite_Union) - (auto simp: disjoint_family_on_def) -qed - -lemma (in finite_prob_space) distribution_eq_setsum_finite: - assumes "finite A" - shows "distribution X A = (\x\A. distribution X {x})" -proof - - note distribution_eq_setsum[of X A] - also have "(\x\A \ X ` space M. distribution X {x}) = (\x\A. distribution X {x})" - proof (intro setsum_mono_zero_cong_left ballI) - fix i assume "i\A - A \ X ` space M" - then have "X -` {i} \ space M = {}" by auto - then show "distribution X {i} = 0" - by (simp add: distribution_def) - next - show "finite A" by fact - qed simp_all + have "(\x. Py x * f x \P) = (\x. f (Y x) \M)" + by (rule distributed_integral) fact+ + also have "\ = (\x. f (T (X x)) \M)" + using Y by simp + also have "\ = (\x. Px x * f (T x) \N)" + using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def) finally show ?thesis . 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), measure = ereal \ 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 - fix A assume "A \ X ` space M" - then show "distribution X A = (\x\A. distribution X {x})" - by (subst distribution_eq_setsum) (simp add: Int_absorb2) +lemma distributed_marginal_eq_joint: + assumes T: "sigma_finite_measure T" + assumes S: "sigma_finite_measure S" + assumes Px: "distributed M S X Px" + assumes Py: "distributed M T Y Py" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "AE y in T. Py y = (\\<^isup>+x. Pxy (x, y) \S)" +proof (rule sigma_finite_measure.density_unique[OF T]) + interpret ST: pair_sigma_finite S T using S T unfolding pair_sigma_finite_def by simp + show "Py \ borel_measurable T" "AE y in T. 0 \ Py y" + "(\x. \\<^isup>+ xa. Pxy (xa, x) \S) \ borel_measurable T" "AE y in T. 0 \ \\<^isup>+ x. Pxy (x, y) \S" + using Pxy[THEN distributed_borel_measurable] + by (auto intro!: Py[THEN distributed_borel_measurable] Py[THEN distributed_AE] + ST.positive_integral_snd_measurable' positive_integral_positive) + + show "density T Py = density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)" + proof (rule measure_eqI) + fix A assume A: "A \ sets (density T Py)" + have *: "\x y. x \ space S \ indicator (space S \ A) (x, y) = indicator A y" + by (auto simp: indicator_def) + have "emeasure (density T Py) A = emeasure (distr M T Y) A" + unfolding Py[THEN distributed_distr_eq_density] .. + also have "\ = emeasure (distr M (S \\<^isub>M T) (\x. (X x, Y x))) (space S \ A)" + using A Px Py Pxy + by (subst (1 2) emeasure_distr) + (auto dest: measurable_space distributed_measurable intro!: arg_cong[where f="emeasure M"]) + also have "\ = emeasure (density (S \\<^isub>M T) Pxy) (space S \ A)" + unfolding Pxy[THEN distributed_distr_eq_density] .. + also have "\ = (\\<^isup>+ x. Pxy x * indicator (space S \ A) x \(S \\<^isub>M T))" + using A Pxy by (simp add: emeasure_density distributed_borel_measurable) + also have "\ = (\\<^isup>+y. \\<^isup>+x. Pxy (x, y) * indicator (space S \ A) (x, y) \S \T)" + using A Pxy + by (subst ST.positive_integral_snd_measurable) (simp_all add: emeasure_density distributed_borel_measurable) + also have "\ = (\\<^isup>+y. (\\<^isup>+x. Pxy (x, y) \S) * indicator A y \T)" + using measurable_comp[OF measurable_Pair1[OF measurable_identity] distributed_borel_measurable[OF Pxy]] + using distributed_borel_measurable[OF Pxy] distributed_AE[OF Pxy, THEN ST.AE_pair] + by (subst (asm) ST.AE_commute) (auto intro!: positive_integral_cong_AE positive_integral_multc cong: positive_integral_cong simp: * comp_def) + also have "\ = emeasure (density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)) A" + using A by (intro emeasure_density[symmetric]) (auto intro!: ST.positive_integral_snd_measurable' Pxy[THEN distributed_borel_measurable]) + finally show "emeasure (density T Py) A = emeasure (density T (\x. \\<^isup>+ xa. Pxy (xa, x) \S)) A" . + qed simp qed -lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M), measure = ereal \ distribution X \" - by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def) +lemma (in prob_space) distr_marginal1: + fixes Pxy :: "('b \ 'c) \ real" + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + defines "Px \ \x. real (\\<^isup>+z. Pxy (x, z) \T)" + shows "distributed M S X Px" + unfolding distributed_def +proof safe + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T by default + + have XY: "(\x. (X x, Y x)) \ measurable M (S \\<^isub>M T)" + using Pxy by (rule distributed_measurable) + then show X: "X \ measurable M S" + unfolding measurable_pair_iff by (simp add: comp_def) + from XY have Y: "Y \ measurable M T" + unfolding measurable_pair_iff by (simp add: comp_def) + + from Pxy show borel: "(\x. ereal (Px x)) \ borel_measurable S" + by (auto intro!: ST.positive_integral_fst_measurable borel_measurable_real_of_ereal dest!: distributed_real_measurable simp: Px_def) -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), measure = ereal \ joint_distribution X Y\" - (is "finite_measure_space ?M") -proof (rule finite_measure_spaceI, simp_all) - show "finite (s1 \ s2)" - using assms by auto -next - fix A assume "A \ (s1 \ s2)" - with assms show "joint_distribution X Y A = (\x\A. joint_distribution X Y {x})" - by (intro distribution_eq_setsum_finite) (auto dest: finite_subset) -qed - -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), - measure = ereal \ 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), - measure = ereal \ joint_distribution X Y \" - (is "finite_prob_space ?S") -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def) - 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" - by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) + interpret Pxy: prob_space "distr M (S \\<^isub>M T) (\x. (X x, Y x))" + using XY by (rule prob_space_distr) + have "(\\<^isup>+ x. max 0 (ereal (- Pxy x)) \(S \\<^isub>M T)) = (\\<^isup>+ x. 0 \(S \\<^isub>M T))" + using Pxy + by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_real_measurable distributed_real_AE) + then have Pxy_integrable: "integrable (S \\<^isub>M T) Pxy" + using Pxy Pxy.emeasure_space_1 + by (simp add: integrable_def emeasure_density positive_integral_max_0 distributed_def borel_measurable_ereal_iff cong: positive_integral_cong) + + show "distr M S X = density S Px" + proof (rule measure_eqI) + fix A assume A: "A \ sets (distr M S X)" + with X Y XY have "emeasure (distr M S X) A = emeasure (distr M (S \\<^isub>M T) (\x. (X x, Y x))) (A \ space T)" + by (auto simp add: emeasure_distr + intro!: arg_cong[where f="emeasure M"] dest: measurable_space) + also have "\ = emeasure (density (S \\<^isub>M T) Pxy) (A \ space T)" + using Pxy by (simp add: distributed_def) + also have "\ = \\<^isup>+ x. \\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T \S" + using A borel Pxy + by (simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric] distributed_def) + also have "\ = \\<^isup>+ x. ereal (Px x) * indicator A x \S" + apply (rule positive_integral_cong_AE) + using Pxy[THEN distributed_real_AE, THEN ST.AE_pair] ST.integrable_fst_measurable(1)[OF Pxy_integrable] AE_space + proof eventually_elim + fix x assume "x \ space S" "AE y in T. 0 \ Pxy (x, y)" and i: "integrable T (\y. Pxy (x, y))" + moreover have eq: "\y. y \ space T \ indicator (A \ space T) (x, y) = indicator A x" + by (auto simp: indicator_def) + ultimately have "(\\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T) = + (\\<^isup>+ y. ereal (Pxy (x, y)) \T) * indicator A x" + using Pxy[THEN distributed_real_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong) + also have "(\\<^isup>+ y. ereal (Pxy (x, y)) \T) = Px x" + using i by (simp add: Px_def ereal_real integrable_def positive_integral_positive) + finally show "(\\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \ space T) (x, y) \T) = ereal (Px x) * indicator A x" . + qed + finally show "emeasure (distr M S X) A = emeasure (density S Px) A" + using A borel Pxy by (simp add: emeasure_density) + qed simp + + show "AE x in S. 0 \ ereal (Px x)" + by (simp add: Px_def positive_integral_positive real_of_ereal_pos) qed -subsection "Borel Measure on {0 ..< 1}" - -definition pborel :: "real measure_space" where - "pborel = lborel.restricted_space {0 ..< 1}" - -lemma space_pborel[simp]: - "space pborel = {0 ..< 1}" - unfolding pborel_def by auto - -lemma sets_pborel: - "A \ sets pborel \ A \ sets borel \ A \ { 0 ..< 1}" - unfolding pborel_def by auto +definition + "simple_distributed M X f \ distributed M (count_space (X`space M)) X (\x. ereal (f x)) \ + finite (X`space M)" -lemma in_pborel[intro, simp]: - "A \ {0 ..< 1} \ A \ sets borel \ A \ sets pborel" - unfolding pborel_def by auto - -interpretation pborel: measure_space pborel - using lborel.restricted_measure_space[of "{0 ..< 1}"] - by (simp add: pborel_def) +lemma simple_distributed: + "simple_distributed M X Px \ distributed M (count_space (X`space M)) X Px" + unfolding simple_distributed_def by auto -interpretation pborel: prob_space pborel -proof - show "measure pborel (space pborel) = 1" - by (simp add: one_ereal_def pborel_def) -qed default - -lemma pborel_prob: "pborel.prob A = (if A \ sets borel \ A \ {0 ..< 1} then real (lborel.\ A) else 0)" - unfolding pborel.\'_def by (auto simp: pborel_def) +lemma simple_distributed_finite[dest]: "simple_distributed M X P \ finite (X`space M)" + by (simp add: simple_distributed_def) -lemma pborel_singelton[simp]: "pborel.prob {a} = 0" - by (auto simp: pborel_prob) - -lemma - shows pborel_atLeastAtMost[simp]: "pborel.\' {a .. b} = (if 0 \ a \ a \ b \ b < 1 then b - a else 0)" - and pborel_atLeastLessThan[simp]: "pborel.\' {a ..< b} = (if 0 \ a \ a \ b \ b \ 1 then b - a else 0)" - and pborel_greaterThanAtMost[simp]: "pborel.\' {a <.. b} = (if 0 \ a \ a \ b \ b < 1 then b - a else 0)" - and pborel_greaterThanLessThan[simp]: "pborel.\' {a <..< b} = (if 0 \ a \ a \ b \ b \ 1 then b - a else 0)" - unfolding pborel_prob - by (auto simp: atLeastAtMost_subseteq_atLeastLessThan_iff - greaterThanAtMost_subseteq_atLeastLessThan_iff greaterThanLessThan_subseteq_atLeastLessThan_iff) - -lemma pborel_lebesgue_measure: - "A \ sets pborel \ pborel.prob A = real (measure lebesgue A)" - by (simp add: sets_pborel pborel_prob) +lemma (in prob_space) distributed_simple_function_superset: + assumes X: "simple_function M X" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" + assumes A: "X`space M \ A" "finite A" + defines "S \ count_space A" and "P' \ (\x. if x \ X`space M then P x else 0)" + shows "distributed M S X P'" + unfolding distributed_def +proof safe + show "(\x. ereal (P' x)) \ borel_measurable S" unfolding S_def by simp + show "AE x in S. 0 \ ereal (P' x)" + using X by (auto simp: S_def P'_def simple_distributed_def intro!: measure_nonneg) + show "distr M S X = density S P'" + proof (rule measure_eqI_finite) + show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A" + using A unfolding S_def by auto + show "finite A" by fact + fix a assume a: "a \ A" + then have "a \ X`space M \ X -` {a} \ space M = {}" by auto + with A a X have "emeasure (distr M S X) {a} = P' a" + by (subst emeasure_distr) + (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure + intro!: arg_cong[where f=prob]) + also have "\ = (\\<^isup>+x. ereal (P' a) * indicator {a} x \S)" + using A X a + by (subst positive_integral_cmult_indicator) + (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg) + also have "\ = (\\<^isup>+x. ereal (P' x) * indicator {a} x \S)" + by (auto simp: indicator_def intro!: positive_integral_cong) + also have "\ = emeasure (density S P') {a}" + using a A by (intro emeasure_density[symmetric]) (auto simp: S_def) + finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" . + qed + show "random_variable S X" + using X(1) A by (auto simp: measurable_def simple_functionD S_def) +qed -lemma pborel_alt: - "pborel = sigma \ - space = {0..<1}, - sets = range (\(x,y). {x.. {0..<1}), - measure = measure lborel \" (is "_ = ?R") +lemma (in prob_space) simple_distributedI: + assumes X: "simple_function M X" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" + shows "simple_distributed M X P" + unfolding simple_distributed_def +proof + have "distributed M (count_space (X ` space M)) X (\x. ereal (if x \ X`space M then P x else 0))" + (is "?A") + using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X]) auto + also have "?A \ distributed M (count_space (X ` space M)) X (\x. ereal (P x))" + by (rule distributed_cong_density) auto + finally show "\" . +qed (rule simple_functionD[OF X(1)]) + +lemma simple_distributed_joint_finite: + assumes X: "simple_distributed M (\x. (X x, Y x)) Px" + shows "finite (X ` space M)" "finite (Y ` space M)" proof - - have *: "{0..<1::real} \ sets borel" by auto - have **: "op \ {0..<1::real} ` range (\(x, y). {x..(x,y). {x.. {0..<1})" - unfolding image_image by (intro arg_cong[where f=range]) auto - have "pborel = algebra.restricted_space (sigma \space=UNIV, sets=range (\(a, b). {a ..< b :: real}), - measure = measure pborel\) {0 ..< 1}" - by (simp add: sigma_def lebesgue_def pborel_def borel_eq_atLeastLessThan lborel_def) - also have "\ = ?R" - by (subst restricted_sigma) - (simp_all add: sets_sigma sigma_sets.Basic ** pborel_def image_eqI[of _ _ "(0,1)"]) - finally show ?thesis . + have "finite ((\x. (X x, Y x)) ` space M)" + using X by (auto simp: simple_distributed_def simple_functionD) + then have "finite (fst ` (\x. (X x, Y x)) ` space M)" "finite (snd ` (\x. (X x, Y x)) ` space M)" + by auto + then show fin: "finite (X ` space M)" "finite (Y ` space M)" + by (auto simp: image_image) +qed + +lemma simple_distributed_joint2_finite: + assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" + shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" +proof - + have "finite ((\x. (X x, Y x, Z x)) ` space M)" + using X by (auto simp: simple_distributed_def simple_functionD) + then have "finite (fst ` (\x. (X x, Y x, Z x)) ` space M)" + "finite ((fst \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" + "finite ((snd \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" + by auto + then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" + by (auto simp: image_image) qed -subsection "Bernoulli space" +lemma simple_distributed_simple_function: + "simple_distributed M X Px \ simple_function M X" + unfolding simple_distributed_def distributed_def + by (auto simp: simple_function_def) + +lemma simple_distributed_measure: + "simple_distributed M X P \ a \ X`space M \ P a = measure M (X -` {a} \ space M)" + using distributed_count_space[of M "X`space M" X P a, symmetric] + by (auto simp: simple_distributed_def measure_def) + +lemma simple_distributed_nonneg: "simple_distributed M X f \ x \ space M \ 0 \ f (X x)" + by (auto simp: simple_distributed_measure measure_nonneg) -definition "bernoulli_space p = \ space = UNIV, sets = UNIV, - measure = ereal \ setsum (\b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \" +lemma (in prob_space) simple_distributed_joint: + assumes X: "simple_distributed M (\x. (X x, Y x)) Px" + defines "S \ count_space (X`space M) \\<^isub>M count_space (Y`space M)" + defines "P \ (\x. if x \ (\x. (X x, Y x))`space M then Px x else 0)" + shows "distributed M S (\x. (X x, Y x)) P" +proof - + from simple_distributed_joint_finite[OF X, simp] + have S_eq: "S = count_space (X`space M \ Y`space M)" + by (simp add: S_def pair_measure_count_space) + show ?thesis + unfolding S_eq P_def + proof (rule distributed_simple_function_superset) + show "simple_function M (\x. (X x, Y x))" + using X by (rule simple_distributed_simple_function) + fix x assume "x \ (\x. (X x, Y x)) ` space M" + from simple_distributed_measure[OF X this] + show "Px x = prob ((\x. (X x, Y x)) -` {x} \ space M)" . + qed auto +qed -interpretation bernoulli: finite_prob_space "bernoulli_space p" for p - by (rule finite_prob_spaceI) - (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg) +lemma (in prob_space) simple_distributed_joint2: + assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" + defines "S \ count_space (X`space M) \\<^isub>M count_space (Y`space M) \\<^isub>M count_space (Z`space M)" + defines "P \ (\x. if x \ (\x. (X x, Y x, Z x))`space M then Px x else 0)" + shows "distributed M S (\x. (X x, Y x, Z x)) P" +proof - + from simple_distributed_joint2_finite[OF X, simp] + have S_eq: "S = count_space (X`space M \ Y`space M \ Z`space M)" + by (simp add: S_def pair_measure_count_space) + show ?thesis + unfolding S_eq P_def + proof (rule distributed_simple_function_superset) + show "simple_function M (\x. (X x, Y x, Z x))" + using X by (rule simple_distributed_simple_function) + fix x assume "x \ (\x. (X x, Y x, Z x)) ` space M" + from simple_distributed_measure[OF X this] + show "Px x = prob ((\x. (X x, Y x, Z x)) -` {x} \ space M)" . + qed auto +qed + +lemma (in prob_space) simple_distributed_setsum_space: + assumes X: "simple_distributed M X f" + shows "setsum f (X`space M) = 1" +proof - + from X have "setsum f (X`space M) = prob (\i\X`space M. X -` {i} \ space M)" + by (subst finite_measure_finite_Union) + (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD + intro!: setsum_cong arg_cong[where f="prob"]) + also have "\ = prob (space M)" + by (auto intro!: arg_cong[where f=prob]) + finally show ?thesis + using emeasure_space_1 by (simp add: emeasure_eq_measure one_ereal_def) +qed -lemma bernoulli_measure: - "0 \ p \ p \ 1 \ bernoulli.prob p B = (\b\B. if b then p else 1 - p)" - unfolding bernoulli.\'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong) +lemma (in prob_space) distributed_marginal_eq_joint_simple: + assumes Px: "simple_function M X" + assumes Py: "simple_distributed M Y Py" + assumes Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" + assumes y: "y \ Y`space M" + shows "Py y = (\x\X`space M. if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" +proof - + note Px = simple_distributedI[OF Px refl] + have *: "\f A. setsum (\x. max 0 (ereal (f x))) A = ereal (setsum (\x. max 0 (f x)) A)" + by (simp add: setsum_ereal[symmetric] zero_ereal_def) + from distributed_marginal_eq_joint[OF sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite + simple_distributed[OF Px] simple_distributed[OF Py] simple_distributed_joint[OF Pxy], + OF Py[THEN simple_distributed_finite] Px[THEN simple_distributed_finite]] + y Px[THEN simple_distributed_finite] Py[THEN simple_distributed_finite] + Pxy[THEN simple_distributed, THEN distributed_real_AE] + show ?thesis + unfolding AE_count_space + apply (elim ballE[where x=y]) + apply (auto simp add: positive_integral_count_space_finite * intro!: setsum_cong split: split_max) + done +qed -lemma bernoulli_measure_True: "0 \ p \ p \ 1 \ bernoulli.prob p {True} = p" - and bernoulli_measure_False: "0 \ p \ p \ 1 \ bernoulli.prob p {False} = 1 - p" - unfolding bernoulli_measure by simp_all + +lemma prob_space_uniform_measure: + assumes A: "emeasure M A \ 0" "emeasure M A \ \" + shows "prob_space (uniform_measure M A)" +proof + show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1" + using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"] + using sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A + by (simp add: Int_absorb2 emeasure_nonneg) +qed + +lemma prob_space_uniform_count_measure: "finite A \ A \ {} \ prob_space (uniform_count_measure A)" + by default (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ereal_def) end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon Apr 23 12:14:35 2012 +0200 @@ -8,45 +8,79 @@ imports Lebesgue_Integration begin +definition "diff_measure M N = + measure_of (space M) (sets M) (\A. emeasure M A - emeasure N A)" + +lemma + shows space_diff_measure[simp]: "space (diff_measure M N) = space M" + and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M" + by (auto simp: diff_measure_def) + +lemma emeasure_diff_measure: + assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N" + assumes pos: "\A. A \ sets M \ emeasure N A \ emeasure M A" and A: "A \ sets M" + shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\ A") + unfolding diff_measure_def +proof (rule emeasure_measure_of_sigma) + show "sigma_algebra (space M) (sets M)" .. + show "positive (sets M) ?\" + using pos by (simp add: positive_def ereal_diff_positive) + show "countably_additive (sets M) ?\" + proof (rule countably_additiveI) + fix A :: "nat \ _" assume A: "range A \ sets M" and "disjoint_family A" + then have suminf: + "(\i. emeasure M (A i)) = emeasure M (\i. A i)" + "(\i. emeasure N (A i)) = emeasure N (\i. A i)" + by (simp_all add: suminf_emeasure sets_eq) + with A have "(\i. emeasure M (A i) - emeasure N (A i)) = + (\i. emeasure M (A i)) - (\i. emeasure N (A i))" + using fin + by (intro suminf_ereal_minus pos emeasure_nonneg) + (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure) + then show "(\i. emeasure M (A i) - emeasure N (A i)) = + emeasure M (\i. A i) - emeasure N (\i. A i) " + by (simp add: suminf) + qed +qed fact + lemma (in sigma_finite_measure) Ex_finite_integrable_function: shows "\h\borel_measurable M. integral\<^isup>P M h \ \ \ (\x\space M. 0 < h x \ h x < \) \ (\x. 0 \ h x)" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and - measure: "\i. \ (A i) \ \" and + measure: "\i. emeasure M (A i) \ \" and disjoint: "disjoint_family A" - using disjoint_sigma_finite by auto - let ?B = "\i. 2^Suc i * \ (A i)" + using sigma_finite_disjoint by auto + let ?B = "\i. 2^Suc i * emeasure M (A i)" have "\i. \x. 0 < x \ x < inverse (?B i)" proof - fix i have Ai: "A i \ sets M" using range by auto - from measure positive_measure[OF this] - show "\x. 0 < x \ x < inverse (?B i)" - by (auto intro!: ereal_dense simp: ereal_0_gt_inverse) + fix i show "\x. 0 < x \ x < inverse (?B i)" + using measure[of i] emeasure_nonneg[of M "A i"] + by (auto intro!: ereal_dense simp: ereal_0_gt_inverse ereal_zero_le_0_iff) qed from choice[OF this] obtain n where n: "\i. 0 < n i" - "\i. n i < inverse (2^Suc i * \ (A i))" by auto + "\i. n i < inverse (2^Suc i * emeasure M (A i))" by auto { fix i have "0 \ n i" using n(1)[of i] by auto } note pos = this let ?h = "\x. \i. n i * indicator (A i) x" show ?thesis proof (safe intro!: bexI[of _ ?h] del: notI) have "\i. A i \ sets M" using range by fastforce+ - then have "integral\<^isup>P M ?h = (\i. n i * \ (A i))" using pos + then have "integral\<^isup>P M ?h = (\i. n i * emeasure M (A i))" using pos by (simp add: positive_integral_suminf positive_integral_cmult_indicator) also have "\ \ (\i. (1 / 2)^Suc i)" proof (rule suminf_le_pos) fix N - have "n N * \ (A N) \ inverse (2^Suc N * \ (A N)) * \ (A N)" - using positive_measure[OF `A N \ sets M`] n[of N] + have "n N * emeasure M (A N) \ inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)" + using n[of N] by (intro ereal_mult_right_mono) auto also have "\ \ (1 / 2) ^ Suc N" using measure[of N] n[of N] - by (cases rule: ereal2_cases[of "n N" "\ (A N)"]) + by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"]) (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide) - finally show "n N * \ (A N) \ (1 / 2) ^ Suc N" . - show "0 \ n N * \ (A N)" using n[of N] `A N \ sets M` by simp + finally show "n N * emeasure M (A N) \ (1 / 2) ^ Suc N" . + show "0 \ n N * emeasure M (A N)" using n[of N] `A N \ sets M` by (simp add: emeasure_nonneg) qed finally show "integral\<^isup>P M ?h \ \" unfolding suminf_half_series_ereal by auto next @@ -71,68 +105,45 @@ subsection "Absolutely continuous" -definition (in measure_space) - "absolutely_continuous \ = (\N\null_sets. \ N = (0 :: ereal))" +definition absolutely_continuous :: "'a measure \ 'a measure \ bool" where + "absolutely_continuous M N \ null_sets M \ null_sets N" + +lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M" + unfolding absolutely_continuous_def by (auto simp: null_sets_count_space) -lemma (in measure_space) absolutely_continuous_AE: - assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M" - and "absolutely_continuous (measure M')" "AE x. P x" +lemma absolutely_continuousI_density: + "f \ borel_measurable M \ absolutely_continuous M (density M f)" + by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in) + +lemma absolutely_continuousI_point_measure_finite: + "(\x. \ x \ A ; f x \ 0 \ \ g x \ 0) \ absolutely_continuous (point_measure A f) (point_measure A g)" + unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff) + +lemma absolutely_continuous_AE: + assumes sets_eq: "sets M' = sets M" + and "absolutely_continuous M M'" "AE x in M. P x" shows "AE x in M'. P x" proof - - 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 + from `AE x in M. P x` obtain N where N: "N \ null_sets M" "{x\space M. \ P x} \ N" + unfolding eventually_ae_filter by auto show "AE x in M'. P x" - proof (rule \.AE_I') - 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 + proof (rule AE_I') + show "{x\space M'. \ P x} \ N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp + from `absolutely_continuous M M'` show "N \ null_sets M'" + using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto qed qed -lemma (in finite_measure_space) absolutely_continuousI: - 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 ?\ by fact - have "\ N = measure ?\ (\x\N. {x})" by simp - also have "\ = (\x\N. measure ?\ {x})" - proof (rule v.measure_setsum[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 ?\" 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" "0 \ \ {x}" - using sets_eq_Pow `N \ space M` positive_measure[of "{x}"] - by (auto intro!: measure_mono) - then have "\ {x} = 0" using `\ N = 0` by simp - thus "measure ?\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto - qed - 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 \M)" - shows "absolutely_continuous \" - using assms unfolding absolutely_continuous_def - by (simp add: positive_integral_null_set) - subsection "Existence of the Radon-Nikodym derivative" lemma (in finite_measure) Radon_Nikodym_aux_epsilon: fixes e :: real assumes "0 < e" - assumes "finite_measure (M\measure := \\)" - shows "\A\sets M. \' (space M) - finite_measure.\' (M\measure := \\) (space M) \ - \' A - finite_measure.\' (M\measure := \\) A \ - (\B\sets M. B \ A \ - e < \' B - finite_measure.\' (M\measure := \\) B)" + assumes "finite_measure N" and sets_eq: "sets N = sets M" + shows "\A\sets M. measure M (space M) - measure N (space M) \ measure M A - measure N A \ + (\B\sets M. B \ A \ - e < measure M B - measure N B)" proof - - interpret M': finite_measure "M\measure := \\" by fact - let ?d = "\A. \' A - M'.\' A" + interpret M': finite_measure N by fact + let ?d = "\A. measure M A - measure N A" 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)" @@ -159,7 +170,7 @@ fix B assume "B \ sets M \ B \ space M - A n \ ?d B \ - e" hence "A n \ B = {}" "B \ sets M" and dB: "?d B \ -e" by auto hence "?d (A n \ B) = ?d (A n) + ?d B" - using `A n \ sets M` finite_measure_Union M'.finite_measure_Union by simp + using `A n \ sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq) also have "\ \ ?d (A n) - e" using dB by simp finally show "?d (A n \ B) \ ?d (A n) - e" . qed } @@ -187,8 +198,8 @@ proof (induct n) fix n assume "?d (space M) \ ?d (space M - A n)" also have "\ \ ?d (space M - A (Suc n))" - using A_in_sets sets_into_space dA_mono[of n] - by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff) + using A_in_sets sets_into_space dA_mono[of n] finite_measure_compl M'.finite_measure_compl + by (simp del: A_simps add: sets_eq sets_eq_imp_space_eq[OF sets_eq]) finally show "?d (space M) \ ?d (space M - A (Suc n))" . qed simp qed @@ -199,17 +210,17 @@ proof (induct n) case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps) next - case 0 with M'.empty_measure show ?case by (simp add: zero_ereal_def) + case 0 with measure_empty show ?case by (simp add: zero_ereal_def) qed } note dA_less = this have decseq: "decseq (\n. ?d (A n))" unfolding decseq_eq_incseq proof (rule incseq_SucI) fix n show "- ?d (A n) \ - ?d (A (Suc n))" using dA_mono[of n] by auto qed have A: "incseq A" by (auto intro!: incseq_SucI) - from finite_continuity_from_below[OF _ A] `range A \ sets M` - M'.finite_continuity_from_below[OF _ A] + from finite_Lim_measure_incseq[OF _ A] `range A \ sets M` + M'.finite_Lim_measure_incseq[OF _ A] have convergent: "(\i. ?d (A i)) ----> ?d (\i. A i)" - by (auto intro!: tendsto_diff) + by (auto intro!: tendsto_diff simp: sets_eq) obtain n :: nat where "- ?d (\i. A i) / e < real n" using reals_Archimedean2 by auto moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less] have "real n \ - ?d (\i. A i) / e" using `0 sets M" and X: "X \ S" - shows "finite_measure.\' (restricted_space S) X = \' X" -proof cases - note r = restricted_finite_measure[OF S] - { assume "X \ sets M" with S X show ?thesis - unfolding finite_measure.\'_def[OF r] \'_def by auto } - { assume "X \ sets M" - moreover then have "S \ X \ sets M" - using X by (simp add: Int_absorb1) - ultimately show ?thesis - unfolding finite_measure.\'_def[OF r] \'_def using S by auto } -qed - -lemma (in finite_measure) restricted_measure: - assumes X: "S \ sets M" "X \ sets (restricted_space S)" - shows "finite_measure.\' (restricted_space S) X = \' X" +lemma (in finite_measure) Radon_Nikodym_aux: + assumes "finite_measure N" and sets_eq: "sets N = sets M" + shows "\A\sets M. measure M (space M) - measure N (space M) \ + measure M A - measure N A \ + (\B\sets M. B \ A \ 0 \ measure M B - measure N B)" proof - - from X have "S \ sets M" "X \ S" by auto - from restricted_measure_subset[OF this] show ?thesis . -qed - -lemma (in finite_measure) Radon_Nikodym_aux: - assumes "finite_measure (M\measure := \\)" (is "finite_measure ?M'") - shows "\A\sets M. \' (space M) - finite_measure.\' (M\measure := \\) (space M) \ - \' A - finite_measure.\' (M\measure := \\) A \ - (\B\sets M. B \ A \ 0 \ \' B - finite_measure.\' (M\measure := \\) B)" -proof - - interpret M': finite_measure ?M' where - "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \" by fact auto - let ?d = "\A. \' A - M'.\' A" + interpret N: finite_measure N by fact + let ?d = "\A. measure M A - measure N 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)" let ?r = "\S. restricted_space S" { fix S n assume S: "S \ sets M" - note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S - then have "finite_measure (?r S)" "0 < 1 / real (Suc n)" - "finite_measure (?r S\measure := \\)" by auto + then have "finite_measure (density M (indicator S))" "0 < 1 / real (Suc n)" + "finite_measure (density N (indicator S))" "sets (density N (indicator S)) = sets (density M (indicator S))" + by (auto simp: finite_measure_restricted N.finite_measure_restricted sets_eq) from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this - have "?P X S n" - proof (intro conjI ballI impI) - show "X \ sets M" "X \ S" using X(1) `S \ sets M` by auto - have "S \ op \ S ` sets M" using `S \ sets M` by auto - then show "?d S \ ?d X" - using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp - fix C assume "C \ sets M" "C \ X" - then have "C \ sets (restricted_space S)" "C \ X" - using `S \ sets M` `X \ S` by auto - with X(2) show "- 1 / real (Suc n) < ?d C" - using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto - qed - hence "\A. ?P A S n" by auto } + with S have "?P (S \ X) S n" + by (simp add: measure_restricted sets_eq Int) (metis inf_absorb2) + hence "\A. ?P A S n" .. } note Ex_P = this def A \ "nat_rec (space M) (\n A. SOME B. ?P B A n)" have A_Suc: "\n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def) @@ -292,12 +270,11 @@ proof (safe intro!: bexI[of _ "\i. A i"]) show "(\i. A i) \ sets M" using A_in_sets by auto have A: "decseq A" using A_mono by (auto intro!: decseq_SucI) - from - finite_continuity_from_above[OF `range A \ sets M` A] - M'.finite_continuity_from_above[OF `range A \ sets M` A] - have "(\i. ?d (A i)) ----> ?d (\i. A i)" by (intro tendsto_diff) + from `range A \ sets M` + finite_Lim_measure_decseq[OF _ A] N.finite_Lim_measure_decseq[OF _ A] + have "(\i. ?d (A i)) ----> ?d (\i. A i)" by (auto intro!: tendsto_diff simp: sets_eq) thus "?d (space M) \ ?d (\i. A i)" using mono_dA[THEN monoD, of 0 _] - by (rule_tac LIMSEQ_le_const) (auto intro!: exI) + by (rule_tac LIMSEQ_le_const) auto next fix B assume B: "B \ sets M" "B \ (\i. A i)" show "0 \ ?d B" @@ -315,14 +292,12 @@ qed lemma (in finite_measure) Radon_Nikodym_finite_measure: - 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 \M)" + assumes "finite_measure N" and sets_eq: "sets N = sets M" + assumes "absolutely_continuous M N" + shows "\f \ borel_measurable M. (\x. 0 \ f x) \ density M f = N" proof - - 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. (\x. 0 \ g x) \ (\A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ \ A)}" + interpret N: finite_measure N by fact + def G \ "{g \ borel_measurable M. (\x. 0 \ g x) \ (\A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ N 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" @@ -333,6 +308,7 @@ have "?A \ sets M" using f g unfolding G_def by auto fix A assume "A \ sets M" hence sets: "?A \ A \ sets M" "(space M - ?A) \ A \ sets M" using `?A \ sets M` by auto + hence sets': "?A \ A \ sets N" "(space M - ?A) \ A \ sets N" by (auto simp: sets_eq) have union: "((?A \ A) \ ((space M - ?A) \ A)) = A" using sets_into_space[OF `A \ sets M`] by auto have "\x. x \ space M \ max (g x) (f x) * indicator A x = @@ -343,11 +319,11 @@ (\\<^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) - also have "\ \ \ (?A \ A) + \ ((space M - ?A) \ A)" + also have "\ \ N (?A \ A) + N ((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 \M) \ \ A" . + also have "\ = N A" + using plus_emeasure[OF sets'] union by auto + finally show "(\\<^isup>+x. max (g x) (f x) * indicator A x \M) \ N A" . next fix x show "0 \ max (g x) (f x)" using f g by (auto simp: G_def split: split_max) qed } @@ -368,15 +344,15 @@ using `incseq f` f `A \ sets M` by (intro positive_integral_monotone_convergence_SUP) (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) - finally show "(\\<^isup>+x. (SUP i. f i x) * indicator A x \M) \ \ A" + finally show "(\\<^isup>+x. (SUP i. f i x) * indicator A x \M) \ N A" using f `A \ sets M` by (auto intro!: SUP_least simp: G_def) qed } note SUP_in_G = this let ?y = "SUP g : G. integral\<^isup>P M g" - have "?y \ \ (space M)" unfolding G_def + have y_le: "?y \ N (space M)" unfolding G_def proof (safe intro!: SUP_least) - 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)" + fix g assume "\A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ N A" + from this[THEN bspec, OF top] show "integral\<^isup>P M g \ N (space M)" by (simp cong: positive_integral_cong) qed from SUPR_countable_SUPR[OF `G \ {}`, of "integral\<^isup>P M"] guess ys .. note ys = this @@ -411,8 +387,7 @@ also have "\ = ?y" proof (rule antisym) show "(SUP i. integral\<^isup>P M (?g i)) \ ?y" - using g_in_G - by (auto intro!: exI Sup_mono simp: SUP_def) + using g_in_G by (auto intro: Sup_mono simp: SUP_def) show "?y \ (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq by (auto intro!: SUP_mono positive_integral_mono Max_ge) qed @@ -420,107 +395,68 @@ have "\x. 0 \ f x" unfolding f_def using `\i. gs i \ G` by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def) - let ?t = "\A. \ A - (\\<^isup>+x. ?F A x \M)" - let ?M = "M\ measure := ?t\" - interpret M: sigma_algebra ?M - by (intro sigma_algebra_cong) auto - have f_le_\: "\A. A \ sets M \ (\\<^isup>+x. ?F A x \M) \ \ A" + let ?t = "\A. N A - (\\<^isup>+x. ?F A x \M)" + let ?M = "diff_measure N (density M f)" + have f_le_N: "\A. A \ sets M \ (\\<^isup>+x. ?F A x \M) \ N A" using `f \ G` unfolding G_def by auto - have fmM: "finite_measure ?M" + have emeasure_M: "\A. A \ sets M \ emeasure ?M A = ?t A" + proof (subst emeasure_diff_measure) + from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)" + by (auto intro!: finite_measureI simp: emeasure_density cong: positive_integral_cong) + next + fix B assume "B \ sets N" with f_le_N[of B] show "emeasure (density M f) B \ emeasure N B" + by (auto simp: sets_eq emeasure_density cong: positive_integral_cong) + qed (auto simp: sets_eq emeasure_density) + from emeasure_M[of "space M"] N.finite_emeasure_space positive_integral_positive[of M "?F (space M)"] + interpret M': finite_measure ?M + by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure ereal_minus_eq_PInfty_iff) + + have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def proof - show "measure_space ?M" - proof (default, simp_all add: countably_additive_def positive_def, safe del: notI) - fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" - have "(\n. (\\<^isup>+x. ?F (A n) x \M)) = (\\<^isup>+x. (\n. ?F (A n) x) \M)" - using `range A \ sets M` `\x. 0 \ f x` - by (intro positive_integral_suminf[symmetric]) auto - also have "\ = (\\<^isup>+x. ?F (\n. A n) x \M)" - using `\x. 0 \ f x` - by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`]) - finally have "(\n. (\\<^isup>+x. ?F (A n) x \M)) = (\\<^isup>+x. ?F (\n. A n) x \M)" . - moreover have "(\n. \ (A n)) = \ (\n. A n)" - using M'.measure_countably_additive A by (simp add: comp_def) - moreover have v_fin: "\ (\i. A i) \ \" using M'.finite_measure A by (simp add: countable_UN) - moreover { - have "(\\<^isup>+x. ?F (\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 - finally have "(\\<^isup>+x. ?F (\i. A i) x \M) \ \" by simp } - moreover have "\i. (\\<^isup>+x. ?F (A i) x \M) \ \ (A i)" - using A by (intro f_le_\) auto - ultimately - show "(\n. ?t (A n)) = ?t (\i. A i)" - by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive) - next - fix A assume A: "A \ sets M" show "0 \ \ A - \\<^isup>+ x. ?F A x \M" - using f_le_\[OF A] `f \ G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff) - qed - next - show "measure ?M (space ?M) \ \" - using positive_integral_positive[of "?F (space M)"] - by (cases rule: ereal2_cases[of "\ (space M)" "\\<^isup>+ x. ?F (space M) x \M"]) auto + fix A assume A: "A \ null_sets M" + with `absolutely_continuous M N` have "A \ null_sets N" + unfolding absolutely_continuous_def by auto + moreover with A have "(\\<^isup>+ x. ?F A x \M) \ N A" using `f \ G` by (auto simp: G_def) + ultimately have "N A - (\\<^isup>+ x. ?F A x \M) = 0" + using positive_integral_positive[of M] by (auto intro!: antisym) + then show "A \ null_sets ?M" + using A by (simp add: emeasure_M null_sets_def sets_eq) qed - 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" unfolding absolutely_continuous_def - proof - fix N assume N: "N \ null_sets" - with `absolutely_continuous \` have "\ N = 0" unfolding absolutely_continuous_def by auto - moreover with N have "(\\<^isup>+ x. ?F N x \M) \ \ N" using `f \ G` by (auto simp: G_def) - ultimately show "\ N - (\\<^isup>+ x. ?F N x \M) = 0" - using positive_integral_positive by (auto intro!: antisym) - qed - have upper_bound: "\A\sets M. ?t A \ 0" + have upper_bound: "\A\sets M. ?M A \ 0" proof (rule ccontr) assume "\ ?thesis" - then obtain A where A: "A \ sets M" and pos: "0 < ?t A" + then obtain A where A: "A \ sets M" and pos: "0 < ?M A" by (auto simp: not_le) note pos - also have "?t A \ ?t (space M)" - using M.measure_mono[of A "space M"] A sets_into_space by simp - finally have pos_t: "0 < ?t (space M)" by simp + also have "?M A \ ?M (space M)" + using emeasure_space[of ?M A] by (simp add: sets_eq[THEN sets_eq_imp_space_eq]) + finally have pos_t: "0 < ?M (space M)" by simp moreover - then have "\ (space M) \ 0" - using ac unfolding absolutely_continuous_def by auto - then have pos_M: "0 < \ (space M)" - using positive_measure[OF top] by (simp add: le_less) + then have "emeasure M (space M) \ 0" + using ac unfolding absolutely_continuous_def by (auto simp: null_sets_def) + then have pos_M: "0 < emeasure M (space M)" + using emeasure_nonneg[of M "space M"] by (simp add: le_less) moreover - have "(\\<^isup>+x. f x * indicator (space M) x \M) \ \ (space M)" + have "(\\<^isup>+x. f x * indicator (space M) x \M) \ N (space M)" using `f \ G` unfolding G_def by auto hence "(\\<^isup>+x. f x * indicator (space M) x \M) \ \" - using M'.finite_measure_of_space by auto + using M'.finite_emeasure_space by auto moreover - def b \ "?t (space M) / \ (space M) / 2" + def b \ "?M (space M) / emeasure M (space M) / 2" ultimately have b: "b \ 0 \ 0 \ b \ b \ \" - using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"] - by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\ (space M)" "\ (space M)"]) - (simp_all add: field_simps) + by (auto simp: ereal_divide_eq) then have b: "b \ 0" "0 \ b" "0 < b" "b \ \" by auto - let ?Mb = "?M\measure := \A. b * \ A\" - interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto - have Mb: "finite_measure ?Mb" - proof - show "measure_space ?Mb" - proof - show "positive ?Mb (measure ?Mb)" - using `0 \ b` by (auto simp: positive_def) - show "countably_additive ?Mb (measure ?Mb)" - using `0 \ b` measure_countably_additive - by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq) - qed - show "measure ?Mb (space ?Mb) \ \" - using b by auto - qed - 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 - *: "\B. \ B \ sets M ; B \ A0 \ \ 0 \ real (?t B) - real (b * \ B)" - unfolding M.\'_def finite_measure.\'_def[OF Mb] by auto + let ?Mb = "density M (\_. b)" + have Mb: "finite_measure ?Mb" "sets ?Mb = sets ?M" + using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI) + from M'.Radon_Nikodym_aux[OF this] guess A0 .. + then have "A0 \ sets M" + and space_less_A0: "measure ?M (space M) - real b * measure M (space M) \ measure ?M A0 - real b * measure M A0" + and *: "\B. B \ sets M \ B \ A0 \ 0 \ measure ?M B - real b * measure M B" + using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq) { fix B assume B: "B \ sets M" "B \ A0" - with *[OF this] have "b * \ B \ ?t B" - using M'.finite_measure b finite_measure M.positive_measure[OF B(1)] - by (cases rule: ereal2_cases[of "?t B" "b * \ B"]) auto } + with *[OF this] have "b * emeasure M B \ ?M B" + using b unfolding M'.emeasure_eq_measure emeasure_eq_measure by (cases b) auto } note bM_le_t = this let ?f0 = "\x. f x + b * indicator A0 x" { fix A assume A: "A \ sets M" @@ -529,71 +465,47 @@ (\\<^isup>+x. f x * indicator A x + b * indicator (A \ A0) x \M)" by (auto intro!: positive_integral_cong split: split_indicator) hence "(\\<^isup>+x. ?f0 x * indicator A x \M) = - (\\<^isup>+x. f x * indicator A x \M) + b * \ (A \ A0)" + (\\<^isup>+x. f x * indicator A x \M) + b * emeasure M (A \ A0)" using `A0 \ sets M` `A \ A0 \ sets M` A b `f \ G` by (simp add: G_def 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 \M) \ \ A" - using `f \ G` A unfolding G_def by auto + have f_le_v: "(\\<^isup>+x. ?F A x \M) \ N A" using `f \ G` A unfolding G_def by auto note f0_eq[OF A] - also have "(\\<^isup>+x. f x * indicator A x \M) + b * \ (A \ A0) \ - (\\<^isup>+x. f x * indicator A x \M) + ?t (A \ A0)" + also have "(\\<^isup>+x. ?F A x \M) + b * emeasure M (A \ A0) \ (\\<^isup>+x. ?F A x \M) + ?M (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 \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`] positive_integral_positive[of "?F A"] - by (cases "\\<^isup>+x. ?F A x \M", cases "\ A") auto - finally have "(\\<^isup>+x. ?f0 x * indicator A x \M) \ \ A" . } + also have "\ \ (\\<^isup>+x. f x * indicator A x \M) + ?M A" + using emeasure_mono[of "A \ A0" A ?M] `A \ sets M` `A0 \ sets M` + by (auto intro!: add_left_mono simp: sets_eq) + also have "\ \ N A" + unfolding emeasure_M[OF `A \ sets M`] + using f_le_v N.emeasure_eq_measure[of A] positive_integral_positive[of M "?F A"] + by (cases "\\<^isup>+x. ?F A x \M", cases "N A") auto + finally have "(\\<^isup>+x. ?f0 x * indicator A x \M) \ N A" . } hence "?f0 \ G" using `A0 \ sets M` b `f \ G` unfolding G_def by (auto intro!: ereal_add_nonneg_nonneg) - have real: "?t (space M) \ \" "?t A0 \ \" - "b * \ (space M) \ \" "b * \ A0 \ \" - using `A0 \ sets M` b - finite_measure[of A0] M.finite_measure[of A0] - finite_measure_of_space M.finite_measure_of_space - by auto have int_f_finite: "integral\<^isup>P M f \ \" - using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff - by (auto cong: positive_integral_cong) - have "0 < ?t (space M) - b * \ (space M)" unfolding b_def - using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M - using positive_integral_positive[of "?F (space M)"] - by (cases rule: ereal3_cases[of "\ (space M)" "\ (space M)" "integral\<^isup>P M (?F (space M))"]) - (auto simp: field_simps mult_less_cancel_left) - also have "\ \ ?t A0 - b * \ A0" - using space_less_A0 b - using - `A0 \ sets M`[THEN M.real_measure] - top[THEN M.real_measure] - apply safe - apply simp - using - `A0 \ sets M`[THEN real_measure] - `A0 \ sets M`[THEN M'.real_measure] - top[THEN real_measure] - top[THEN M'.real_measure] - by (cases b) auto - finally have 1: "b * \ A0 < ?t A0" - using - `A0 \ sets M`[THEN M.real_measure] - apply safe - apply simp - using - `A0 \ sets M`[THEN real_measure] - `A0 \ sets M`[THEN M'.real_measure] - by (cases b) auto - have "0 < ?t A0" - using b `A0 \ sets M` by (auto intro!: le_less_trans[OF _ 1]) - then have "\ A0 \ 0" using ac unfolding absolutely_continuous_def - using `A0 \ sets M` by auto - then have "0 < \ A0" using positive_measure[OF `A0 \ sets M`] by auto - hence "0 < b * \ A0" using b by (auto simp: ereal_zero_less_0_iff) - with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \ A0" unfolding int_f_eq_y + by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le) + have "0 < ?M (space M) - emeasure ?Mb (space M)" + using pos_t + by (simp add: b emeasure_density_const) + (simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def) + also have "\ \ ?M A0 - b * emeasure M A0" + using space_less_A0 `A0 \ sets M` b + by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure) + finally have 1: "b * emeasure M A0 < ?M A0" + by (metis M'.emeasure_real `A0 \ sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1) + less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def) + with b have "0 < ?M A0" + by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times + ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def) + then have "emeasure M A0 \ 0" using ac `A0 \ sets M` + by (auto simp: absolutely_continuous_def null_sets_def) + then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto + hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff) + with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * emeasure M A0" unfolding int_f_eq_y using `f \ G` by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive) also have "\ = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \ sets M` sets_into_space @@ -602,75 +514,67 @@ moreover from `?f0 \ G` have "integral\<^isup>P M ?f0 \ ?y" by (auto intro!: SUP_upper) ultimately show False by auto qed + let ?f = "\x. max 0 (f x)" show ?thesis - proof (safe intro!: bexI[of _ f]) - fix A assume A: "A\sets M" - show "\ A = (\\<^isup>+x. f x * indicator A x \M)" - proof (rule antisym) - 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 \M)" - using upper_bound[THEN bspec, OF `A \ sets M`] - using M'.real_measure[OF A] - by (cases "integral\<^isup>P M (?F A)") auto - qed - qed simp + proof (intro bexI[of _ ?f] measure_eqI conjI) + show "sets (density M ?f) = sets N" + by (simp add: sets_eq) + fix A assume A: "A\sets (density M ?f)" + then show "emeasure (density M ?f) A = emeasure N A" + using `f \ G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A] + by (cases "integral\<^isup>P M (?F A)") + (auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric]) + qed auto qed lemma (in finite_measure) split_space_into_finite_sets_and_rest: - assumes "measure_space (M\measure := \\)" (is "measure_space ?N") - assumes ac: "absolutely_continuous \" + assumes ac: "absolutely_continuous M N" and sets_eq: "sets N = sets M" shows "\A0\sets M. \B::nat\'a set. disjoint_family B \ range B \ sets M \ A0 = space M - (\i. B i) \ - (\A\sets M. A \ A0 \ (\ A = 0 \ \ A = 0) \ (\ A > 0 \ \ A = \)) \ - (\i. \ (B i) \ \)" + (\A\sets M. A \ A0 \ (emeasure M A = 0 \ N A = 0) \ (emeasure M A > 0 \ N A = \)) \ + (\i. N (B i) \ \)" proof - - 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 + let ?Q = "{Q\sets M. N Q \ \}" + let ?a = "SUP Q:?Q. emeasure M Q" + have "{} \ ?Q" by auto then have Q_not_empty: "?Q \ {}" by blast - have "?a \ \ (space M)" using sets_into_space - by (auto intro!: SUP_least measure_mono) - then have "?a \ \" using finite_measure_of_space + have "?a \ emeasure M (space M)" using sets_into_space + by (auto intro!: SUP_least emeasure_mono) + then have "?a \ \" using finite_emeasure_space by auto - from SUPR_countable_SUPR[OF Q_not_empty, of \] - obtain Q'' where "range Q'' \ \ ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" + from SUPR_countable_SUPR[OF Q_not_empty, of "emeasure M"] + obtain Q'' where "range Q'' \ emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" by auto - then have "\i. \Q'. Q'' i = \ Q' \ Q' \ ?Q" by auto - from choice[OF this] obtain Q' where Q': "\i. Q'' i = \ (Q' i)" "\i. Q' i \ ?Q" + then have "\i. \Q'. Q'' i = emeasure M Q' \ Q' \ ?Q" by auto + from choice[OF this] obtain Q' where Q': "\i. Q'' i = emeasure M (Q' i)" "\i. Q' i \ ?Q" by auto - then have a_Lim: "?a = (SUP i::nat. \ (Q' i))" using a by simp + then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp let ?O = "\n. \i\n. Q' i" - have Union: "(SUP i. \ (?O i)) = \ (\i. ?O i)" - proof (rule continuity_from_below[of ?O]) - show "range ?O \ sets M" using Q' by (auto intro!: finite_UN) + have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\i. ?O i)" + proof (rule SUP_emeasure_incseq[of ?O]) + show "range ?O \ sets M" using Q' by auto show "incseq ?O" by (fastforce intro!: incseq_SucI) qed have Q'_sets: "\i. Q' i \ sets M" using Q' by auto - have O_sets: "\i. ?O i \ sets M" - using Q' by (auto intro!: finite_UN Un) + have O_sets: "\i. ?O i \ sets M" using Q' by auto then have O_in_G: "\i. ?O i \ ?Q" proof (safe del: notI) - fix i have "Q' ` {..i} \ sets M" - using Q' by (auto intro: finite_UN) - with v.measure_finitely_subadditive[of "{.. i}" Q'] - have "\ (?O i) \ (\i\i. \ (Q' i))" by auto + fix i have "Q' ` {..i} \ sets M" using Q' by auto + then have "N (?O i) \ (\i\i. N (Q' i))" + by (simp add: sets_eq emeasure_subadditive_finite) also have "\ < \" using Q' by (simp add: setsum_Pinfty) - finally show "\ (?O i) \ \" by simp + finally show "N (?O i) \ \" by simp qed auto have O_mono: "\n. ?O n \ ?O (Suc n)" by fastforce - have a_eq: "?a = \ (\i. ?O i)" unfolding Union[symmetric] + have a_eq: "?a = emeasure M (\i. ?O i)" unfolding Union[symmetric] proof (rule antisym) - show "?a \ (SUP i. \ (?O i))" unfolding a_Lim - using Q' by (auto intro!: SUP_mono measure_mono finite_UN) - show "(SUP i. \ (?O i)) \ ?a" unfolding SUP_def + show "?a \ (SUP i. emeasure M (?O i))" unfolding a_Lim + using Q' by (auto intro!: SUP_mono emeasure_mono) + show "(SUP i. emeasure M (?O i)) \ ?a" unfolding SUP_def proof (safe intro!: Sup_mono, unfold bex_simps) fix i have *: "(\Q' ` {..i}) = ?O i" by auto - then show "\x. (x \ sets M \ \ x \ \) \ - \ (\Q' ` {..i}) \ \ x" + then show "\x. (x \ sets M \ N x \ \) \ + emeasure M (\Q' ` {..i}) \ emeasure M x" using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) qed qed @@ -687,51 +591,50 @@ show "range Q \ sets M" using Q_sets by auto { fix A assume A: "A \ sets M" "A \ space M - ?O_0" - show "\ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" + show "emeasure M A = 0 \ N A = 0 \ 0 < emeasure M A \ N A = \" proof (rule disjCI, simp) - assume *: "0 < \ A \ \ A \ \" - show "\ A = 0 \ \ A = 0" + assume *: "0 < emeasure M A \ N A \ \" + show "emeasure M A = 0 \ N A = 0" proof cases - assume "\ A = 0" moreover with ac A have "\ A = 0" + assume "emeasure M A = 0" moreover with ac A have "N A = 0" unfolding absolutely_continuous_def by auto ultimately show ?thesis by simp next - assume "\ A \ 0" with * have "\ A \ \" using positive_measure[OF A(1)] by auto - with A have "\ ?O_0 + \ A = \ (?O_0 \ A)" - using Q' by (auto intro!: measure_additive countable_UN) - also have "\ = (SUP i. \ (?O i \ A))" - proof (rule continuity_from_below[of "\i. ?O i \ A", symmetric, simplified]) + assume "emeasure M A \ 0" with * have "N A \ \" using emeasure_nonneg[of M A] by auto + with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \ A)" + using Q' by (auto intro!: plus_emeasure countable_UN) + also have "\ = (SUP i. emeasure M (?O i \ A))" + proof (rule SUP_emeasure_incseq[of "\i. ?O i \ A", symmetric, simplified]) show "range (\i. ?O i \ A) \ sets M" - using `\ A \ \` O_sets A by auto + using `N A \ \` O_sets A by auto qed (fastforce intro!: incseq_SucI) also have "\ \ ?a" proof (safe intro!: SUP_least) fix i have "?O i \ A \ ?Q" proof (safe del: notI) show "?O i \ A \ sets M" using O_sets A by auto - from O_in_G[of i] have "\ (?O i \ A) \ \ (?O i) + \ A" - using v.measure_subadditive[of "?O i" A] A O_sets by auto - with O_in_G[of i] show "\ (?O i \ A) \ \" - using `\ A \ \` by auto + from O_in_G[of i] have "N (?O i \ A) \ N (?O i) + N A" + using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq) + with O_in_G[of i] show "N (?O i \ A) \ \" + using `N A \ \` by auto qed - then show "\ (?O i \ A) \ ?a" by (rule SUP_upper) + then show "emeasure M (?O i \ A) \ ?a" by (rule SUP_upper) qed - finally have "\ A = 0" - unfolding a_eq using real_measure[OF `?O_0 \ sets M`] real_measure[OF A(1)] by auto - with `\ A \ 0` show ?thesis by auto + finally have "emeasure M A = 0" + unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure) + with `emeasure M A \ 0` show ?thesis by auto qed qed } - { fix i show "\ (Q i) \ \" + { fix i show "N (Q i) \ \" proof (cases i) case 0 then show ?thesis unfolding Q_def using Q'[of 0] by simp next case (Suc n) - then show ?thesis unfolding Q_def - using `?O n \ ?Q` `?O (Suc n) \ ?Q` - using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"] - using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono] - by (cases rule: ereal2_cases[of "\ (\ x\Suc n. Q' x)" "\ (\ i\n. Q' i)"]) auto + with `?O n \ ?Q` `?O (Suc n) \ ?Q` + emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\ x\n. Q' x)"] + show ?thesis + by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def) qed } show "space M - ?O_0 \ sets M" using Q'_sets by auto { fix j have "(\i\j. ?O i) = (\i\j. Q i)" @@ -751,96 +654,93 @@ qed lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite: - assumes "measure_space (M\measure := \\)" (is "measure_space ?N") - assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. (\x. 0 \ f x) \ (\A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" + assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M" + shows "\f\borel_measurable M. (\x. 0 \ f x) \ density M f = N" proof - - 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" and Q0: "Q0 \ sets M" "Q0 = space M - (\i. Q i)" - and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" - and Q_fin: "\i. \ (Q i) \ \" by force + and in_Q0: "\A. A \ sets M \ A \ Q0 \ emeasure M A = 0 \ N A = 0 \ 0 < emeasure M A \ N A = \" + and Q_fin: "\i. N (Q i) \ \" by force from Q have Q_sets: "\i. Q i \ sets M" by auto - have "\i. \f. f\borel_measurable M \ (\x. 0 \ f x) \ (\A\sets M. - \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" - proof + let ?N = "\i. density N (indicator (Q i))" and ?M = "\i. density M (indicator (Q i))" + have "\i. \f\borel_measurable (?M i). (\x. 0 \ f x) \ density (?M i) f = ?N i" + proof (intro allI finite_measure.Radon_Nikodym_finite_measure) fix i - have indicator_eq: "\f x A. (f x :: ereal) * 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]]) - then interpret R: finite_measure ?R . - have fmv: "finite_measure (restricted_space (Q i) \ measure := \\)" (is "finite_measure ?Q") - proof - show "measure_space ?Q" - using v.restricted_measure_space Q_sets[of i] by auto - 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 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 \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 \ (\x. 0 \ f x) \ (\A\sets M. - \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" - by (auto intro!: exI[of _ "\x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos - split: split_indicator split_if_asm simp: max_def) + from Q show "finite_measure (?M i)" + by (auto intro!: finite_measureI cong: positive_integral_cong + simp add: emeasure_density subset_eq sets_eq) + from Q have "emeasure (?N i) (space N) = emeasure N (Q i)" + by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: positive_integral_cong) + with Q_fin show "finite_measure (?N i)" + by (auto intro!: finite_measureI) + show "sets (?N i) = sets (?M i)" by (simp add: sets_eq) + show "absolutely_continuous (?M i) (?N i)" + using `absolutely_continuous M N` `Q i \ sets M` + by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq + intro!: absolutely_continuous_AE[OF sets_eq]) qed - from choice[OF this] obtain f where borel: "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" - and f: "\A i. A \ sets M \ - \ (Q i \ A) = (\\<^isup>+x. f i x * indicator (Q i \ A) x \M)" + from choice[OF this[unfolded Bex_def]] + obtain f where borel: "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + and f_density: "\i. density (?M i) (f i) = ?N i" by auto + { fix A i assume A: "A \ sets M" + with Q borel have "(\\<^isup>+x. f i x * indicator (Q i \ A) x \M) = emeasure (density (?M i) (f i)) A" + by (auto simp add: emeasure_density positive_integral_density subset_eq + intro!: positive_integral_cong split: split_indicator) + also have "\ = emeasure N (Q i \ A)" + using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq) + finally have "emeasure N (Q i \ A) = (\\<^isup>+x. f i x * indicator (Q i \ A) x \M)" .. } + note integral_eq = this let ?f = "\x. (\i. f i x * indicator (Q i) x) + \ * indicator Q0 x" show ?thesis proof (safe intro!: bexI[of _ ?f]) show "?f \ borel_measurable M" using Q0 borel Q_sets by (auto intro!: measurable_If) show "\x. 0 \ ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def) - fix A assume "A \ sets M" - have Qi: "\i. Q i \ sets M" using Q by auto - have [intro,simp]: "\i. (\x. f i x * indicator (Q i \ A) x) \ borel_measurable M" - "\i. AE x. 0 \ f i x * indicator (Q i \ A) x" - using borel Qi Q0(1) `A \ sets M` by (auto intro!: borel_measurable_ereal_times) - have "(\\<^isup>+x. ?f x * indicator A x \M) = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) + \ * indicator (Q0 \ A) x \M)" - using borel by (intro positive_integral_cong) (auto simp: indicator_def) - also have "\ = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) \M) + \ * \ (Q0 \ A)" - using borel Qi Q0(1) `A \ sets M` - by (subst positive_integral_add) (auto simp del: ereal_infty_mult - simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le) - also have "\ = (\i. \ (Q i \ A)) + \ * \ (Q0 \ A)" - by (subst f[OF `A \ sets M`], subst positive_integral_suminf) auto - finally have "(\\<^isup>+x. ?f x * indicator A x \M) = (\i. \ (Q i \ A)) + \ * \ (Q0 \ A)" . - moreover have "(\i. \ (Q i \ A)) = \ ((\i. Q i) \ A)" - using Q Q_sets `A \ sets M` - by (intro v.measure_countably_additive[of "\i. Q i \ A", unfolded comp_def, simplified]) - (auto simp: disjoint_family_on_def) - moreover have "\ * \ (Q0 \ A) = \ (Q0 \ A)" - proof - - have "Q0 \ A \ sets M" using Q0(1) `A \ sets M` by blast - from in_Q0[OF this] show ?thesis by auto - qed - moreover have "Q0 \ A \ sets M" "((\i. Q i) \ A) \ sets M" - 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 \M)" - using v.measure_additive[simplified, of "(\i. Q i) \ A" "Q0 \ A"] - by simp + show "density M ?f = N" + proof (rule measure_eqI) + fix A assume "A \ sets (density M ?f)" + then have "A \ sets M" by simp + have Qi: "\i. Q i \ sets M" using Q by auto + have [intro,simp]: "\i. (\x. f i x * indicator (Q i \ A) x) \ borel_measurable M" + "\i. AE x in M. 0 \ f i x * indicator (Q i \ A) x" + using borel Qi Q0(1) `A \ sets M` by (auto intro!: borel_measurable_ereal_times) + have "(\\<^isup>+x. ?f x * indicator A x \M) = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) + \ * indicator (Q0 \ A) x \M)" + using borel by (intro positive_integral_cong) (auto simp: indicator_def) + also have "\ = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) \M) + \ * emeasure M (Q0 \ A)" + using borel Qi Q0(1) `A \ sets M` + by (subst positive_integral_add) (auto simp del: ereal_infty_mult + simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le) + also have "\ = (\i. N (Q i \ A)) + \ * emeasure M (Q0 \ A)" + by (subst integral_eq[OF `A \ sets M`], subst positive_integral_suminf) auto + finally have "(\\<^isup>+x. ?f x * indicator A x \M) = (\i. N (Q i \ A)) + \ * emeasure M (Q0 \ A)" . + moreover have "(\i. N (Q i \ A)) = N ((\i. Q i) \ A)" + using Q Q_sets `A \ sets M` + by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq) + moreover have "\ * emeasure M (Q0 \ A) = N (Q0 \ A)" + proof - + have "Q0 \ A \ sets M" using Q0(1) `A \ sets M` by blast + from in_Q0[OF this] show ?thesis by auto + qed + moreover have "Q0 \ A \ sets M" "((\i. Q i) \ A) \ sets M" + using Q_sets `A \ sets M` Q0(1) by auto + 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 have "N A = (\\<^isup>+x. ?f x * indicator A x \M)" + using plus_emeasure[of "(\i. Q i) \ A" N "Q0 \ A"] by (simp add: sets_eq) + with `A \ sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A" + by (subst emeasure_density) + (auto intro!: borel_measurable_ereal_add borel_measurable_psuminf measurable_If + simp: subset_eq) + qed (simp add: sets_eq) qed qed lemma (in sigma_finite_measure) Radon_Nikodym: - assumes "measure_space (M\measure := \\)" (is "measure_space ?N") - assumes ac: "absolutely_continuous \" - shows "\f \ borel_measurable M. (\x. 0 \ f x) \ (\A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" + assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M" + shows "\f \ borel_measurable M. (\x. 0 \ f x) \ density M f = N" proof - from Ex_finite_integrable_function obtain h where finite: "integral\<^isup>P M h \ \" and @@ -849,47 +749,38 @@ 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 \M)" - let ?MT = "M\ measure := ?T \" - interpret T: finite_measure ?MT - where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T" - using borel finite nn - by (auto intro!: measure_space_density finite_measureI cong: positive_integral_cong) - have "T.absolutely_continuous \" - proof (unfold T.absolutely_continuous_def, safe) - fix N assume "N \ sets M" "(\\<^isup>+x. h x * indicator N x \M) = 0" - with borel ac pos have "AE x. x \ N" - by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le) - then have "N \ null_sets" using `N \ sets M` sets_into_space - by (subst (asm) AE_iff_measurable[OF `N \ sets M`]) auto - then show "\ N = 0" using ac by (auto simp: absolutely_continuous_def) - qed - from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this] - obtain f where f_borel: "f \ borel_measurable M" "\x. 0 \ f x" and - 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" - using borel f_borel by auto - show "\x. 0 \ h x * f x" using nn f_borel by auto - fix A assume "A \ sets M" - then show "\ A = (\\<^isup>+x. h x * f x * indicator A x \M)" - unfolding fT[OF `A \ sets M`] mult_assoc using nn borel f_borel - by (intro positive_integral_translated_density) auto - qed + let ?MT = "density M h" + from borel finite nn interpret T: finite_measure ?MT + by (auto intro!: finite_measureI cong: positive_integral_cong simp: emeasure_density) + have "absolutely_continuous ?MT N" "sets N = sets ?MT" + proof (unfold absolutely_continuous_def, safe) + fix A assume "A \ null_sets ?MT" + with borel have "A \ sets M" "AE x in M. x \ A \ h x \ 0" + by (auto simp add: null_sets_density_iff) + with pos sets_into_space have "AE x in M. x \ A" + by (elim eventually_elim1) (auto simp: not_le[symmetric]) + then have "A \ null_sets M" + using `A \ sets M` by (simp add: AE_iff_null_sets) + with ac show "A \ null_sets N" + by (auto simp: absolutely_continuous_def) + qed (auto simp add: sets_eq) + from T.Radon_Nikodym_finite_measure_infinite[OF this] + obtain f where f_borel: "f \ borel_measurable M" "\x. 0 \ f x" "density ?MT f = N" by auto + with nn borel show ?thesis + by (auto intro!: bexI[of _ "\x. h x * f x"] simp: density_density_eq) qed section "Uniqueness of densities" -lemma (in measure_space) finite_density_unique: +lemma finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" - assumes pos: "AE x. 0 \ f x" "AE x. 0 \ g x" + assumes pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ g 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)" + \ (AE x in M. f x = g x)" (is "(\A\sets M. ?P f A = ?P g A) \ _") proof (intro iffI ballI) - fix A assume eq: "AE x. f x = g x" + fix A assume eq: "AE x in M. f x = g x" then show "?P f A = ?P g A" by (auto intro: positive_integral_cong_AE) next @@ -897,7 +788,7 @@ from this[THEN bspec, OF top] fin 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 pos: "AE x. 0 \ f x" "AE x. 0 \ g x" + and pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ g x" 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 @@ -910,167 +801,163 @@ 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 - show "AE x. g x * indicator ?N x \ f x * indicator ?N x" - "AE x. 0 \ g x * indicator ?N x" + show "AE x in M. g x * indicator ?N x \ f x * indicator ?N x" + "AE x in M. 0 \ g x * indicator ?N x" using pos by (auto split: split_indicator) qed fact also have "\ = 0" - unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto - finally have "AE x. f x \ g x" + unfolding eq[THEN bspec, OF N] using positive_integral_positive[of M] Pg_fin by auto + finally have "AE x in M. f x \ g x" using pos borel positive_integral_PInf_AE[OF borel(2) g_fin] by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_less ereal_minus_le_iff) } from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq - show "AE x. f x = g x" by auto + show "AE x in M. f x = g x" by auto qed lemma (in finite_measure) density_unique_finite_measure: assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" - assumes pos: "AE x. 0 \ f x" "AE x. 0 \ f' x" + assumes pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ f' 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" + shows "AE x in M. f x = f' x" proof - - let ?\ = "\A. ?P f A" and ?\' = "\A. ?P f' A" + let ?D = "\f. density M f" + let ?N = "\A. ?P f A" and ?N' = "\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\ measure := ?\\" - using borel(1) pos(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\ measure := ?\\)` ac] + + have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M" + using borel by (auto intro!: absolutely_continuousI_density) + from split_space_into_finite_sets_and_rest[OF this] 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)" - and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ ?\ A = 0 \ 0 < \ A \ ?\ A = \" - and Q_fin: "\i. ?\ (Q i) \ \" by force + and in_Q0: "\A. A \ sets M \ A \ Q0 \ emeasure M A = 0 \ ?D f A = 0 \ 0 < emeasure M A \ ?D f A = \" + and Q_fin: "\i. ?D f (Q i) \ \" by force + with borel pos have in_Q0: "\A. A \ sets M \ A \ Q0 \ emeasure M A = 0 \ ?N A = 0 \ 0 < emeasure M A \ ?N A = \" + and Q_fin: "\i. ?N (Q i) \ \" by (auto simp: emeasure_density subset_eq) + from Q have Q_sets: "\i. Q i \ sets M" by auto - let ?N = "{x\space M. f x \ f' x}" - have "?N \ sets M" using borel by auto + let ?D = "{x\space M. f x \ f' x}" + have "?D \ sets M" using borel by auto have *: "\i x A. \y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \ A) x" unfolding indicator_def by auto - have "\i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos + have "\i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos by (intro finite_density_unique[THEN iffD1] allI) (auto intro!: borel_measurable_ereal_times f Int simp: *) - moreover have "AE x. ?f Q0 x = ?f' Q0 x" + moreover have "AE x in M. ?f Q0 x = ?f' Q0 x" proof (rule AE_I') { fix f :: "'a \ ereal" assume borel: "f \ borel_measurable M" - and eq: "\A. A \ sets M \ ?\ A = (\\<^isup>+x. f x * indicator A x \M)" + and eq: "\A. A \ sets M \ ?N A = (\\<^isup>+x. f x * indicator A x \M)" let ?A = "\i. Q0 \ {x \ space M. f x < (i::nat)}" - have "(\i. ?A i) \ null_sets" + have "(\i. ?A i) \ null_sets M" proof (rule null_sets_UN) fix i ::nat have "?A i \ sets M" using borel Q0(1) by auto - have "?\ (?A i) \ (\\<^isup>+x. (i::ereal) * indicator (?A i) x \M)" + have "?N (?A i) \ (\\<^isup>+x. (i::ereal) * indicator (?A i) x \M)" unfolding eq[OF `?A i \ sets M`] by (auto intro!: positive_integral_mono simp: indicator_def) - also have "\ = i * \ (?A i)" + also have "\ = i * emeasure M (?A i)" using `?A i \ sets M` by (auto intro!: positive_integral_cmult_indicator) - also have "\ < \" - using `?A i \ sets M`[THEN finite_measure] by auto - finally have "?\ (?A i) \ \" by simp - then show "?A i \ null_sets" using in_Q0[OF `?A i \ sets M`] `?A i \ sets M` by auto + also have "\ < \" using emeasure_real[of "?A i"] by simp + finally have "?N (?A i) \ \" by simp + then show "?A i \ null_sets M" using in_Q0[OF `?A i \ sets M`] `?A i \ sets M` by auto qed also have "(\i. ?A i) = Q0 \ {x\space M. f x \ \}" by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat) - finally have "Q0 \ {x\space M. f x \ \} \ null_sets" by simp } + finally have "Q0 \ {x\space M. f x \ \} \ null_sets M" by simp } from this[OF borel(1) refl] this[OF borel(2) f] - have "Q0 \ {x\space M. f x \ \} \ null_sets" "Q0 \ {x\space M. f' x \ \} \ null_sets" by simp_all - then show "(Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \}) \ null_sets" by (rule nullsets.Un) + have "Q0 \ {x\space M. f x \ \} \ null_sets M" "Q0 \ {x\space M. f' x \ \} \ null_sets M" by simp_all + then show "(Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \}) \ null_sets M" by (rule null_sets.Un) show "{x \ space M. ?f Q0 x \ ?f' Q0 x} \ (Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \})" by (auto simp: indicator_def) qed - moreover have "\x. (?f Q0 x = ?f' Q0 x) \ (\i. ?f (Q i) x = ?f' (Q i) x) \ + moreover have "AE x in M. (?f Q0 x = ?f' Q0 x) \ (\i. ?f (Q i) x = ?f' (Q i) x) \ ?f (space M) x = ?f' (space M) x" by (auto simp: indicator_def Q0) - ultimately have "AE x. ?f (space M) x = ?f' (space M) x" - by (auto simp: AE_all_countable[symmetric]) - then show "AE x. f x = f' x" by auto + ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x" + unfolding AE_all_countable[symmetric] + by eventually_elim (auto intro!: AE_I2 split: split_if_asm simp: indicator_def) + then show "AE x in M. f x = f' x" by auto qed lemma (in sigma_finite_measure) density_unique: - assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" - assumes f': "f' \ borel_measurable M" "AE x. 0 \ f' x" - assumes eq: "\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" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes f': "f' \ borel_measurable M" "AE x in M. 0 \ f' x" + assumes density_eq: "density M f = density M f'" + shows "AE x in M. f x = f' x" proof - obtain h where h_borel: "h \ borel_measurable M" and fin: "integral\<^isup>P M h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" "\x. 0 \ h x" using Ex_finite_integrable_function by auto - then have h_nn: "AE x. 0 \ h x" by auto - let ?H = "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" - have H: "measure_space ?H" - using h_borel h_nn by (rule measure_space_density) simp - then interpret h: measure_space ?H . - interpret h: finite_measure "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" - by (intro H finite_measureI) (simp cong: positive_integral_cong add: fin) - let ?fM = "M\measure := \A. (\\<^isup>+x. f x * indicator A x \M)\" - interpret f: measure_space ?fM - using f 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 f' by (rule measure_space_density) simp + then have h_nn: "AE x in M. 0 \ h x" by auto + let ?H = "density M h" + interpret h: finite_measure ?H + using fin h_borel pos + by (intro finite_measureI) (simp cong: positive_integral_cong emeasure_density add: fin) + let ?fM = "density M f" + let ?f'M = "density M f'" { fix A assume "A \ sets M" then have "{x \ space M. h x * indicator A x \ 0} = A" using pos(1) sets_into_space by (force simp: indicator_def) - then have "(\\<^isup>+x. h x * indicator A x \M) = 0 \ A \ null_sets" + then have "(\\<^isup>+x. h x * indicator A x \M) = 0 \ A \ null_sets M" using h_borel `A \ sets M` h_nn by (subst positive_integral_0_iff) auto } note h_null_sets = this { fix A assume "A \ sets M" have "(\\<^isup>+x. f x * (h x * indicator A x) \M) = (\\<^isup>+x. h x * indicator A x \?fM)" using `A \ sets M` h_borel h_nn f f' - by (intro positive_integral_translated_density[symmetric]) auto + by (intro positive_integral_density[symmetric]) auto also have "\ = (\\<^isup>+x. h x * indicator A x \?f'M)" - by (rule f'.positive_integral_cong_measure) (simp_all add: eq) + by (simp_all add: density_eq) also have "\ = (\\<^isup>+x. f' x * (h x * indicator A x) \M)" using `A \ sets M` h_borel h_nn f f' - by (intro positive_integral_translated_density) auto + by (intro positive_integral_density) auto finally have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" by (simp add: ac_simps) then have "(\\<^isup>+x. (f x * indicator A x) \?H) = (\\<^isup>+x. (f' x * indicator A x) \?H)" using `A \ sets M` h_borel h_nn f f' - by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto } + by (subst (asm) (1 2) positive_integral_density[symmetric]) auto } then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' - by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous) - simp_all - then show "AE x. f x = f' x" - unfolding h.almost_everywhere_def almost_everywhere_def - by (auto simp add: h_null_sets) + by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) + (auto simp add: AE_density) + then show "AE x in M. f x = f' x" + unfolding eventually_ae_filter using h_borel pos + by (auto simp add: h_null_sets null_sets_density_iff not_less[symmetric] + AE_iff_null_sets[symmetric]) + blast qed +lemma (in sigma_finite_measure) density_unique_iff: + assumes f: "f \ borel_measurable M" and "AE x in M. 0 \ f x" + assumes f': "f' \ borel_measurable M" and "AE x in M. 0 \ f' x" + shows "density M f = density M f' \ (AE x in M. f x = f' x)" + using density_unique[OF assms] density_cong[OF f f'] by auto + lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: - assumes \: "measure_space (M\measure := \\)" (is "measure_space ?N") - and f: "f \ borel_measurable M" "AE x. 0 \ f x" - 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 \ \)" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + shows "sigma_finite_measure (density M f) \ (AE x in M. f x \ \)" + (is "sigma_finite_measure ?N \ _") proof 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 + then interpret N: sigma_finite_measure ?N . + from N.Ex_finite_integrable_function obtain h where h: "h \ borel_measurable M" "integral\<^isup>P ?N h \ \" and h_nn: "\x. 0 \ h x" and fin: "\x\space M. 0 < h x \ h x < \" by auto - have "AE x. f x * h x \ \" + have "AE x in M. f x * h x \ \" proof (rule AE_I') have "integral\<^isup>P ?N h = (\\<^isup>+x. f x * h x \M)" using f h h_nn - by (subst \.positive_integral_cong_measure[symmetric, - of "M\ measure := \ A. \\<^isup>+x. f x * indicator A x \M \"]) - (auto intro!: positive_integral_translated_density simp: eq) + by (auto intro!: positive_integral_density) 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" + then show "(\x. f x * h x) -` {\} \ space M \ null_sets M" using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage) qed auto - then show "AE x. f x \ \" + then show "AE x in M. f x \ \" using fin by (auto elim!: AE_Ball_mp) next - assume AE: "AE x. f x \ \" + assume AE: "AE x in M. f x \ \" from sigma_finite guess Q .. note Q = this - 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 \ {.. ereal(of_nat (Suc n))}) \ space M" { fix i j have "A i \ Q j \ sets M" unfolding A_def using f Q @@ -1113,7 +1000,7 @@ 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" + have "AE x in M. f x * indicator (A i \ Q j) x = 0" using AE by (auto simp: A_def `i = 0`) from positive_integral_cong_AE[OF this] show ?thesis by simp next @@ -1121,271 +1008,238 @@ then have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) \ (\\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \M)" by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat) - also have "\ = Suc n * \ (Q j)" + also have "\ = Suc n * emeasure M (Q j)" using Q by (auto intro!: positive_integral_cmult_indicator) also have "\ < \" using Q by (auto simp: real_eq_of_nat[symmetric]) finally show ?thesis by simp qed - then show "measure ?N (?A n) \ \" - using A_in_sets Q eq by auto + then show "emeasure ?N (?A n) \ \" + using A_in_sets Q f by (auto simp: emeasure_density) qed qed section "Radon-Nikodym derivative" definition - "RN_deriv M \ \ SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ - (\A \ sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" + "RN_deriv M N \ SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N" -lemma (in sigma_finite_measure) RN_deriv_cong: - 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" +lemma + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + shows borel_measurable_RN_deriv_density: "RN_deriv M (density M f) \ borel_measurable M" (is ?borel) + and density_RN_deriv_density: "density M (RN_deriv M (density M f)) = density M f" (is ?density) + and RN_deriv_density_nonneg: "0 \ RN_deriv M (density M f) x" (is ?pos) proof - - interpret \': sigma_finite_measure M' - using cong by (rule sigma_finite_measure_cong) - show ?thesis - unfolding RN_deriv_def - by (simp add: cong \ positive_integral_cong_measure[OF cong] measurable_def) + let ?f = "\x. max 0 (f x)" + let ?P = "\g. g \ borel_measurable M \ (\x. 0 \ g x) \ density M g = density M f" + from f have "?P ?f" using f by (auto intro!: density_cong simp: split: split_max) + then have "?P (RN_deriv M (density M f))" + unfolding RN_deriv_def by (rule someI[where P="?P"]) + then show ?borel ?density ?pos by auto qed lemma (in sigma_finite_measure) RN_deriv: - assumes "measure_space (M\measure := \\)" - assumes "absolutely_continuous \" - 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") - and "0 \ RN_deriv M \ x" + assumes "absolutely_continuous M N" "sets N = sets M" + shows borel_measurable_RN_deriv: "RN_deriv M N \ borel_measurable M" (is ?borel) + and density_RN_deriv: "density M (RN_deriv M N) = N" (is ?density) + and RN_deriv_nonneg: "0 \ RN_deriv M N x" (is ?pos) proof - note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] - then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto - from Ex show "0 \ RN_deriv M \ x" unfolding RN_deriv_def - by (rule someI2_ex) simp - fix A assume "A \ sets M" - from Ex show "?int A" unfolding RN_deriv_def - by (rule someI2_ex) (simp add: `A \ sets M`) + from Ex show ?borel unfolding RN_deriv_def by (rule someI2_ex) simp + from Ex show ?density unfolding RN_deriv_def by (rule someI2_ex) simp + from Ex show ?pos unfolding RN_deriv_def by (rule someI2_ex) simp qed lemma (in sigma_finite_measure) RN_deriv_positive_integral: - assumes \: "measure_space (M\measure := \\)" "absolutely_continuous \" + assumes N: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" - shows "integral\<^isup>P (M\measure := \\) f = (\\<^isup>+x. RN_deriv M \ x * f x \M)" + shows "integral\<^isup>P N f = (\\<^isup>+x. RN_deriv M N x * f x \M)" proof - - interpret \: measure_space "M\measure := \\" by fact - note RN = RN_deriv[OF \] - have "integral\<^isup>P (M\measure := \\) f = (\\<^isup>+x. max 0 (f x) \M\measure := \\)" - unfolding positive_integral_max_0 .. - also have "(\\<^isup>+x. max 0 (f x) \M\measure := \\) = - (\\<^isup>+x. max 0 (f x) \M\measure := \A. (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)\)" - by (intro \.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2)) - also have "\ = (\\<^isup>+x. RN_deriv M \ x * max 0 (f x) \M)" - by (intro positive_integral_translated_density) (auto simp add: RN f) - also have "\ = (\\<^isup>+x. RN_deriv M \ x * f x \M)" - using RN_deriv(3)[OF \] - by (auto intro!: positive_integral_cong_pos split: split_if_asm - simp: max_def ereal_mult_le_0_iff) - finally show ?thesis . -qed - -lemma (in sigma_finite_measure) RN_deriv_unique: - assumes \: "measure_space (M\measure := \\)" "absolutely_continuous \" - and f: "f \ borel_measurable M" "AE x. 0 \ f 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 \]]) - show "AE x. 0 \ RN_deriv M \ x" using RN_deriv[OF \] by auto - fix A assume A: "A \ sets M" - 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] .. + have "integral\<^isup>P N f = integral\<^isup>P (density M (RN_deriv M N)) f" + using N by (simp add: density_RN_deriv) + also have "\ = (\\<^isup>+x. RN_deriv M N x * f x \M)" + using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density) + finally show ?thesis by simp qed -lemma (in sigma_finite_measure) RN_deriv_vimage: - assumes T: "T \ measure_preserving M M'" - and T': "T' \ measure_preserving (M'\ measure := \' \) (M\ measure := \ \)" - and inv: "\x. x \ space M \ T' (T x) = x" - and \': "measure_space (M'\measure := \'\)" "measure_space.absolutely_continuous M' \'" - shows "AE x. RN_deriv M' \' (T x) = RN_deriv M \ x" +lemma null_setsD_AE: "N \ null_sets M \ AE x in M. x \ N" + using AE_iff_null_sets[of N M] by auto + +lemma (in sigma_finite_measure) RN_deriv_unique: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and eq: "density M f = N" + shows "AE x in M. f x = RN_deriv M N x" +proof (rule density_unique) + have ac: "absolutely_continuous M N" + using f(1) unfolding eq[symmetric] by (rule absolutely_continuousI_density) + have eq2: "sets N = sets M" + unfolding eq[symmetric] by simp + show "RN_deriv M N \ borel_measurable M" "AE x in M. 0 \ RN_deriv M N x" + "density M f = density M (RN_deriv M N)" + using RN_deriv[OF ac eq2] eq by auto +qed fact+ + +lemma (in sigma_finite_measure) RN_deriv_distr: + fixes T :: "'a \ 'b" + assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" + and inv: "\x\space M. T' (T x) = x" + and ac: "absolutely_continuous (distr M M' T) (distr N M' T)" + and N: "sets N = sets M" + shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x" proof (rule RN_deriv_unique) - interpret \': measure_space "M'\measure := \'\" by fact - show "measure_space (M\ measure := \\)" - by (rule \'.measure_space_vimage[OF _ T'], simp) default - interpret M': measure_space M' - proof (rule measure_space_vimage) - have "sigma_algebra (M'\ measure := \'\)" by default - then show "sigma_algebra M'" by simp - qed fact - show "absolutely_continuous \" unfolding absolutely_continuous_def - proof safe - fix N assume N: "N \ sets M" and N_0: "\ N = 0" - then have N': "T' -` N \ space M' \ sets M'" - using T' by (auto simp: measurable_def measure_preserving_def) - have "T -` (T' -` N \ space M') \ space M = N" - using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def) - then have "measure M' (T' -` N \ space M') = 0" - using measure_preservingD[OF T N'] N_0 by auto - with \'(2) N' show "\ N = 0" using measure_preservingD[OF T', of N] N - unfolding M'.absolutely_continuous_def measurable_def by auto - qed - interpret M': sigma_finite_measure M' + have [simp]: "sets N = sets M" by fact + note sets_eq_imp_space_eq[OF N, simp] + have measurable_N[simp]: "\M'. measurable N M' = measurable M M'" by (auto simp: measurable_def) + { fix A assume "A \ sets M" + with inv T T' sets_into_space[OF this] + have "T -` T' -` A \ T -` space M' \ space M = A" + by (auto simp: measurable_def) } + note eq = this[simp] + { fix A assume "A \ sets M" + with inv T T' sets_into_space[OF this] + have "(T' \ T) -` A \ space M = A" + by (auto simp: measurable_def) } + note eq2 = this[simp] + let ?M' = "distr M M' T" and ?N' = "distr N M' T" + interpret M': sigma_finite_measure ?M' proof from sigma_finite guess F .. note F = this - show "\A::nat \ 'c set. range A \ sets M' \ (\i. A i) = space M' \ (\i. M'.\ (A i) \ \)" + show "\A::nat \ 'b set. range A \ sets ?M' \ (\i. A i) = space ?M' \ (\i. emeasure ?M' (A i) \ \)" proof (intro exI conjI allI) - show *: "range (\i. T' -` F i \ space M') \ sets M'" - using F T' by (auto simp: measurable_def measure_preserving_def) - show "(\i. T' -` F i \ space M') = space M'" - using F T' by (force simp: measurable_def measure_preserving_def) + show *: "range (\i. T' -` F i \ space ?M') \ sets ?M'" + using F T' by (auto simp: measurable_def) + show "(\i. T' -` F i \ space ?M') = space ?M'" + using F T' by (force simp: measurable_def) fix i have "T' -` F i \ space M' \ sets M'" using * by auto - note measure_preservingD[OF T this, symmetric] moreover have Fi: "F i \ sets M" using F by auto - then have "T -` (T' -` F i \ space M') \ space M = F i" - using T inv sets_into_space[OF Fi] - by (auto simp: measurable_def measure_preserving_def) - ultimately show "measure M' (T' -` F i \ space M') \ \" - using F by simp + ultimately show "emeasure ?M' (T' -` F i \ space ?M') \ \" + using F T T' by (simp add: emeasure_distr) qed qed - have "(RN_deriv M' \') \ T \ borel_measurable M" - by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+ - then show "(\x. RN_deriv M' \' (T x)) \ borel_measurable M" + have "(RN_deriv ?M' ?N') \ T \ borel_measurable M" + using T ac by (intro measurable_comp[where b="?M'"] M'.borel_measurable_RN_deriv) simp_all + then show "(\x. RN_deriv ?M' ?N' (T x)) \ borel_measurable M" by (simp add: comp_def) - show "AE x. 0 \ RN_deriv M' \' (T x)" using M'.RN_deriv(3)[OF \'] by auto - fix A let ?A = "T' -` A \ space M'" - assume A: "A \ sets M" - then have A': "?A \ sets M'" using T' unfolding measurable_def measure_preserving_def - by auto - from A have "\ A = \' ?A" using T'[THEN measure_preservingD] by simp - also have "\ = \\<^isup>+ x. RN_deriv M' \' x * indicator ?A x \M'" - using A' by (rule M'.RN_deriv(2)[OF \']) - also have "\ = \\<^isup>+ x. RN_deriv M' \' (T x) * indicator ?A (T x) \M" - proof (rule positive_integral_vimage) - show "sigma_algebra M'" by default - show "(\x. RN_deriv M' \' x * indicator (T' -` A \ space M') x) \ borel_measurable M'" - by (auto intro!: A' M'.RN_deriv(1)[OF \']) - qed fact - also have "\ = \\<^isup>+ x. RN_deriv M' \' (T x) * indicator A x \M" - using T inv by (auto intro!: positive_integral_cong simp: measure_preserving_def measurable_def indicator_def) - finally show "\ A = \\<^isup>+ x. RN_deriv M' \' (T x) * indicator A x \M" . + show "AE x in M. 0 \ RN_deriv ?M' ?N' (T x)" using M'.RN_deriv_nonneg[OF ac] by auto + + have "N = distr N M (T' \ T)" + by (subst measure_of_of_measure[of N, symmetric]) + (auto simp add: distr_def sigma_sets_eq intro!: measure_of_eq space_closed) + also have "\ = distr (distr N M' T) M T'" + using T T' by (simp add: distr_distr) + also have "\ = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'" + using ac by (simp add: M'.density_RN_deriv) + also have "\ = density M (RN_deriv (distr M M' T) (distr N M' T) \ T)" + using M'.borel_measurable_RN_deriv[OF ac] by (simp add: distr_density_distr[OF T T', OF inv]) + finally show "density M (\x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N" + by (simp add: comp_def) qed lemma (in sigma_finite_measure) RN_deriv_finite: - assumes sfm: "sigma_finite_measure (M\measure := \\)" and ac: "absolutely_continuous \" - shows "AE x. RN_deriv M \ x \ \" + assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" + shows "AE x in M. RN_deriv M N x \ \" proof - - 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(1)[OF \ ac]] RN_deriv(2,3)[OF \ ac] by simp + interpret N: sigma_finite_measure N by fact + from N show ?thesis + using sigma_finite_iff_density_finite[OF RN_deriv(1)[OF ac]] RN_deriv(2,3)[OF ac] by simp qed lemma (in sigma_finite_measure) - assumes \: "sigma_finite_measure (M\measure := \\)" "absolutely_continuous \" + assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" - 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) + shows RN_deriv_integrable: "integrable N f \ + integrable M (\x. real (RN_deriv M N x) * f x)" (is ?integrable) + and RN_deriv_integral: "integral\<^isup>L N f = + (\x. real (RN_deriv M N x) * f x \M)" (is ?integral) proof - - interpret \: sigma_finite_measure "M\measure := \\" by fact - have ms: "measure_space (M\measure := \\)" by default + note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp] + interpret N: sigma_finite_measure N by fact have minus_cong: "\A B A' B'::ereal. 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 - have Nf: "f \ borel_measurable (M\measure := \\)" using f by simp + have Nf: "f \ borel_measurable N" using f by (simp add: measurable_def) { fix f :: "'a \ real" - { fix x assume *: "RN_deriv M \ x \ \" - have "ereal (real (RN_deriv M \ x)) * ereal (f x) = ereal (real (RN_deriv M \ x) * f x)" + { fix x assume *: "RN_deriv M N x \ \" + have "ereal (real (RN_deriv M N x)) * ereal (f x) = ereal (real (RN_deriv M N x) * f x)" by (simp add: mult_le_0_iff) - then have "RN_deriv M \ x * ereal (f x) = ereal (real (RN_deriv M \ x) * f x)" - using RN_deriv(3)[OF ms \(2)] * by (auto simp add: ereal_real split: split_if_asm) } - then have "(\\<^isup>+x. ereal (real (RN_deriv M \ x) * f x) \M) = (\\<^isup>+x. RN_deriv M \ x * ereal (f x) \M)" - "(\\<^isup>+x. ereal (- (real (RN_deriv M \ x) * f x)) \M) = (\\<^isup>+x. RN_deriv M \ x * ereal (- f x) \M)" - using RN_deriv_finite[OF \] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric] + then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)" + using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) } + then have "(\\<^isup>+x. ereal (real (RN_deriv M N x) * f x) \M) = (\\<^isup>+x. RN_deriv M N x * ereal (f x) \M)" + "(\\<^isup>+x. ereal (- (real (RN_deriv M N x) * f x)) \M) = (\\<^isup>+x. RN_deriv M N x * ereal (- f x) \M)" + using RN_deriv_finite[OF N ac] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric] by (auto intro!: positive_integral_cong_AE) } note * = this show ?integral ?integrable unfolding lebesgue_integral_def integrable_def * - using f RN_deriv(1)[OF ms \(2)] - by (auto simp: RN_deriv_positive_integral[OF ms \(2)]) + using Nf f RN_deriv(1)[OF ac] + by (auto simp: RN_deriv_positive_integral[OF ac]) qed lemma (in sigma_finite_measure) real_RN_deriv: - assumes "finite_measure (M\measure := \\)" (is "finite_measure ?\") - assumes ac: "absolutely_continuous \" + assumes "finite_measure N" + assumes ac: "absolutely_continuous M N" "sets N = sets M" obtains D where "D \ borel_measurable M" - and "AE x. RN_deriv M \ x = ereal (D x)" - and "AE x in M\measure := \\. 0 < D x" + and "AE x in M. RN_deriv M N x = ereal (D x)" + and "AE x in N. 0 < D x" and "\x. 0 \ D x" proof - interpret \: finite_measure ?\ by fact - have ms: "measure_space ?\" by default - note RN = RN_deriv[OF ms ac] + interpret N: finite_measure N by fact + + note RN = RN_deriv[OF ac] - let ?RN = "\t. {x \ space M. RN_deriv M \ x = t}" + let ?RN = "\t. {x \ space M. RN_deriv M N x = t}" - show "(\x. real (RN_deriv M \ x)) \ borel_measurable M" + show "(\x. real (RN_deriv M N x)) \ borel_measurable M" using RN by auto - have "\ (?RN \) = (\\<^isup>+ x. RN_deriv M \ x * indicator (?RN \) x \M)" - using RN by auto + have "N (?RN \) = (\\<^isup>+ x. RN_deriv M N x * indicator (?RN \) x \M)" + using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density) also have "\ = (\\<^isup>+ x. \ * indicator (?RN \) x \M)" by (intro positive_integral_cong) (auto simp: indicator_def) - also have "\ = \ * \ (?RN \)" + also have "\ = \ * emeasure M (?RN \)" using RN by (intro positive_integral_cmult_indicator) auto - finally have eq: "\ (?RN \) = \ * \ (?RN \)" . + finally have eq: "N (?RN \) = \ * emeasure M (?RN \)" . moreover - have "\ (?RN \) = 0" + have "emeasure M (?RN \) = 0" proof (rule ccontr) - assume "\ {x \ space M. RN_deriv M \ x = \} \ 0" - moreover from RN have "0 \ \ {x \ space M. RN_deriv M \ x = \}" by auto - ultimately have "0 < \ {x \ space M. RN_deriv M \ x = \}" by auto - with eq have "\ (?RN \) = \" by simp - with \.finite_measure[of "?RN \"] RN show False by auto + assume "emeasure M {x \ space M. RN_deriv M N x = \} \ 0" + moreover from RN have "0 \ emeasure M {x \ space M. RN_deriv M N x = \}" by auto + ultimately have "0 < emeasure M {x \ space M. RN_deriv M N x = \}" by auto + with eq have "N (?RN \) = \" by simp + with N.emeasure_finite[of "?RN \"] RN show False by auto qed - ultimately have "AE x. RN_deriv M \ x < \" + ultimately have "AE x in M. RN_deriv M N x < \" using RN by (intro AE_iff_measurable[THEN iffD2]) auto - then show "AE x. RN_deriv M \ x = ereal (real (RN_deriv M \ x))" + then show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))" using RN(3) by (auto simp: ereal_real) - then have eq: "AE x in (M\measure := \\) . RN_deriv M \ x = ereal (real (RN_deriv M \ x))" - using ac absolutely_continuous_AE[OF ms] by auto + then have eq: "AE x in N. RN_deriv M N x = ereal (real (RN_deriv M N x))" + using ac absolutely_continuous_AE by auto - show "\x. 0 \ real (RN_deriv M \ x)" + show "\x. 0 \ real (RN_deriv M N x)" using RN by (auto intro: real_of_ereal_pos) - have "\ (?RN 0) = (\\<^isup>+ x. RN_deriv M \ x * indicator (?RN 0) x \M)" - using RN by auto + have "N (?RN 0) = (\\<^isup>+ x. RN_deriv M N x * indicator (?RN 0) x \M)" + using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density) also have "\ = (\\<^isup>+ x. 0 \M)" by (intro positive_integral_cong) (auto simp: indicator_def) - finally have "AE x in (M\measure := \\). RN_deriv M \ x \ 0" - using RN by (intro \.AE_iff_measurable[THEN iffD2]) auto - with RN(3) eq show "AE x in (M\measure := \\). 0 < real (RN_deriv M \ x)" + finally have "AE x in N. RN_deriv M N x \ 0" + using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq) + with RN(3) eq show "AE x in N. 0 < real (RN_deriv M N x)" by (auto simp: zero_less_real_of_ereal le_less) qed lemma (in sigma_finite_measure) RN_deriv_singleton: - assumes "measure_space (M\measure := \\)" - and ac: "absolutely_continuous \" - and "{x} \ sets M" - shows "\ {x} = RN_deriv M \ x * \ {x}" + assumes ac: "absolutely_continuous M N" "sets N = sets M" + and x: "{x} \ sets M" + shows "N {x} = RN_deriv M N x * emeasure M {x}" proof - - note deriv = RN_deriv[OF assms(1, 2)] - from deriv(2)[OF `{x} \ sets M`] - 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`] deriv(3) - by auto -qed - -theorem (in finite_measure_space) RN_deriv_finite_measure: - assumes "measure_space (M\measure := \\)" - and ac: "absolutely_continuous \" - and "x \ space M" - 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 . + note deriv = RN_deriv[OF ac] + from deriv(1,3) `{x} \ sets M` + have "density M (RN_deriv M N) {x} = (\\<^isup>+w. RN_deriv M N x * indicator {x} w \M)" + by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong) + with x deriv show ?thesis + by (auto simp: positive_integral_cmult_indicator) qed end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Apr 23 12:14:35 2012 +0200 @@ -13,6 +13,7 @@ "~~/src/HOL/Library/Countable" "~~/src/HOL/Library/FuncSet" "~~/src/HOL/Library/Indicator_Function" + "~~/src/HOL/Library/Extended_Real" begin text {* Sigma algebras are an elementary concept in measure @@ -25,203 +26,198 @@ subsection {* Algebras *} -record 'a algebra = - space :: "'a set" - sets :: "'a set set" +locale subset_class = + fixes \ :: "'a set" and M :: "'a set set" + assumes space_closed: "M \ Pow \" -locale subset_class = - fixes M :: "('a, 'b) algebra_scheme" - assumes space_closed: "sets M \ Pow (space M)" - -lemma (in subset_class) sets_into_space: "x \ sets M \ x \ space M" +lemma (in subset_class) sets_into_space: "x \ M \ x \ \" by (metis PowD contra_subsetD space_closed) locale ring_of_sets = subset_class + - assumes empty_sets [iff]: "{} \ sets M" - and Diff [intro]: "\a b. a \ sets M \ b \ sets M \ a - b \ sets M" - and Un [intro]: "\a b. a \ sets M \ b \ sets M \ a \ b \ sets M" + assumes empty_sets [iff]: "{} \ M" + and Diff [intro]: "\a b. a \ M \ b \ M \ a - b \ M" + and Un [intro]: "\a b. a \ M \ b \ M \ a \ b \ M" lemma (in ring_of_sets) Int [intro]: - assumes a: "a \ sets M" and b: "b \ sets M" shows "a \ b \ sets M" + assumes a: "a \ M" and b: "b \ M" shows "a \ b \ M" proof - have "a \ b = a - (a - b)" by auto - then show "a \ b \ sets M" + then show "a \ b \ M" using a b by auto qed lemma (in ring_of_sets) finite_Union [intro]: - "finite X \ X \ sets M \ Union X \ sets M" + "finite X \ X \ M \ Union X \ M" by (induct set: finite) (auto simp add: Un) lemma (in ring_of_sets) finite_UN[intro]: - assumes "finite I" and "\i. i \ I \ A i \ sets M" - shows "(\i\I. A i) \ sets M" + assumes "finite I" and "\i. i \ I \ A i \ M" + shows "(\i\I. A i) \ M" using assms by induct auto lemma (in ring_of_sets) finite_INT[intro]: - assumes "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" - shows "(\i\I. A i) \ sets M" + assumes "finite I" "I \ {}" "\i. i \ I \ A i \ M" + shows "(\i\I. A i) \ M" using assms by (induct rule: finite_ne_induct) auto lemma (in ring_of_sets) insert_in_sets: - assumes "{x} \ sets M" "A \ sets M" shows "insert x A \ sets M" + assumes "{x} \ M" "A \ M" shows "insert x A \ M" proof - - have "{x} \ A \ sets M" using assms by (rule Un) + have "{x} \ A \ M" using assms by (rule Un) thus ?thesis by auto qed -lemma (in ring_of_sets) Int_space_eq1 [simp]: "x \ sets M \ space M \ x = x" +lemma (in ring_of_sets) Int_space_eq1 [simp]: "x \ M \ \ \ x = x" by (metis Int_absorb1 sets_into_space) -lemma (in ring_of_sets) Int_space_eq2 [simp]: "x \ sets M \ x \ space M = x" +lemma (in ring_of_sets) Int_space_eq2 [simp]: "x \ M \ x \ \ = x" by (metis Int_absorb2 sets_into_space) lemma (in ring_of_sets) sets_Collect_conj: - assumes "{x\space M. P x} \ sets M" "{x\space M. Q x} \ sets M" - shows "{x\space M. Q x \ P x} \ sets M" + assumes "{x\\. P x} \ M" "{x\\. Q x} \ M" + shows "{x\\. Q x \ P x} \ M" proof - - have "{x\space M. Q x \ P x} = {x\space M. Q x} \ {x\space M. P x}" + have "{x\\. Q x \ P x} = {x\\. Q x} \ {x\\. P x}" by auto with assms show ?thesis by auto qed lemma (in ring_of_sets) sets_Collect_disj: - assumes "{x\space M. P x} \ sets M" "{x\space M. Q x} \ sets M" - shows "{x\space M. Q x \ P x} \ sets M" + assumes "{x\\. P x} \ M" "{x\\. Q x} \ M" + shows "{x\\. Q x \ P x} \ M" proof - - have "{x\space M. Q x \ P x} = {x\space M. Q x} \ {x\space M. P x}" + have "{x\\. Q x \ P x} = {x\\. Q x} \ {x\\. P x}" by auto with assms show ?thesis by auto qed lemma (in ring_of_sets) sets_Collect_finite_All: - assumes "\i. i \ S \ {x\space M. P i x} \ sets M" "finite S" "S \ {}" - shows "{x\space M. \i\S. P i x} \ sets M" + assumes "\i. i \ S \ {x\\. P i x} \ M" "finite S" "S \ {}" + shows "{x\\. \i\S. P i x} \ M" proof - - have "{x\space M. \i\S. P i x} = (\i\S. {x\space M. P i x})" + have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" using `S \ {}` by auto with assms show ?thesis by auto qed lemma (in ring_of_sets) sets_Collect_finite_Ex: - assumes "\i. i \ S \ {x\space M. P i x} \ sets M" "finite S" - shows "{x\space M. \i\S. P i x} \ sets M" + assumes "\i. i \ S \ {x\\. P i x} \ M" "finite S" + shows "{x\\. \i\S. P i x} \ M" proof - - have "{x\space M. \i\S. P i x} = (\i\S. {x\space M. P i x})" + have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" by auto with assms show ?thesis by auto qed locale algebra = ring_of_sets + - assumes top [iff]: "space M \ sets M" + assumes top [iff]: "\ \ M" lemma (in algebra) compl_sets [intro]: - "a \ sets M \ space M - a \ sets M" + "a \ M \ \ - a \ M" by auto lemma algebra_iff_Un: - "algebra M \ - sets M \ Pow (space M) & - {} \ sets M & - (\a \ sets M. space M - a \ sets M) & - (\a \ sets M. \ b \ sets M. a \ b \ sets M)" (is "_ \ ?Un") + "algebra \ M \ + M \ Pow \ \ + {} \ M \ + (\a \ M. \ - a \ M) \ + (\a \ M. \ b \ M. a \ b \ M)" (is "_ \ ?Un") proof - assume "algebra M" - then interpret algebra M . + assume "algebra \ M" + then interpret algebra \ M . show ?Un using sets_into_space by auto next assume ?Un - show "algebra M" + show "algebra \ M" proof - show space: "sets M \ Pow (space M)" "{} \ sets M" "space M \ sets M" + show \: "M \ Pow \" "{} \ M" "\ \ M" using `?Un` by auto - fix a b assume a: "a \ sets M" and b: "b \ sets M" - then show "a \ b \ sets M" using `?Un` by auto - have "a - b = space M - ((space M - a) \ b)" - using space a b by auto - then show "a - b \ sets M" + fix a b assume a: "a \ M" and b: "b \ M" + then show "a \ b \ M" using `?Un` by auto + have "a - b = \ - ((\ - a) \ b)" + using \ a b by auto + then show "a - b \ M" using a b `?Un` by auto qed qed lemma algebra_iff_Int: - "algebra M \ - sets M \ Pow (space M) & {} \ sets M & - (\a \ sets M. space M - a \ sets M) & - (\a \ sets M. \ b \ sets M. a \ b \ sets M)" (is "_ \ ?Int") + "algebra \ M \ + M \ Pow \ & {} \ M & + (\a \ M. \ - a \ M) & + (\a \ M. \ b \ M. a \ b \ M)" (is "_ \ ?Int") proof - assume "algebra M" - then interpret algebra M . + assume "algebra \ M" + then interpret algebra \ M . show ?Int using sets_into_space by auto next assume ?Int - show "algebra M" + show "algebra \ M" proof (unfold algebra_iff_Un, intro conjI ballI) - show space: "sets M \ Pow (space M)" "{} \ sets M" + show \: "M \ Pow \" "{} \ M" using `?Int` by auto - from `?Int` show "\a. a \ sets M \ space M - a \ sets M" by auto - fix a b assume sets: "a \ sets M" "b \ sets M" - hence "a \ b = space M - ((space M - a) \ (space M - b))" - using space by blast - also have "... \ sets M" - using sets `?Int` by auto - finally show "a \ b \ sets M" . + from `?Int` show "\a. a \ M \ \ - a \ M" by auto + fix a b assume M: "a \ M" "b \ M" + hence "a \ b = \ - ((\ - a) \ (\ - b))" + using \ by blast + also have "... \ M" + using M `?Int` by auto + finally show "a \ b \ M" . qed qed lemma (in algebra) sets_Collect_neg: - assumes "{x\space M. P x} \ sets M" - shows "{x\space M. \ P x} \ sets M" + assumes "{x\\. P x} \ M" + shows "{x\\. \ P x} \ M" proof - - have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto + have "{x\\. \ P x} = \ - {x\\. P x}" by auto with assms show ?thesis by auto qed lemma (in algebra) sets_Collect_imp: - "{x\space M. P x} \ sets M \ {x\space M. Q x} \ sets M \ {x\space M. Q x \ P x} \ sets M" + "{x\\. P x} \ M \ {x\\. Q x} \ M \ {x\\. Q x \ P x} \ M" unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg) lemma (in algebra) sets_Collect_const: - "{x\space M. P} \ sets M" + "{x\\. P} \ M" by (cases P) auto lemma algebra_single_set: assumes "X \ S" - shows "algebra \ space = S, sets = { {}, X, S - X, S }\" + shows "algebra S { {}, X, S - X, S }" by default (insert `X \ S`, auto) section {* Restricted algebras *} abbreviation (in algebra) - "restricted_space A \ \ space = A, sets = (\S. (A \ S)) ` sets M, \ = more M \" + "restricted_space A \ (op \ A) ` M" lemma (in algebra) restricted_algebra: - assumes "A \ sets M" shows "algebra (restricted_space A)" + assumes "A \ M" shows "algebra A (restricted_space A)" using assms by unfold_locales auto subsection {* Sigma Algebras *} locale sigma_algebra = algebra + - assumes countable_nat_UN [intro]: - "!!A. range A \ sets M \ (\i::nat. A i) \ sets M" + assumes countable_nat_UN [intro]: "\A. range A \ M \ (\i::nat. A i) \ M" lemma (in algebra) is_sigma_algebra: - assumes "finite (sets M)" - shows "sigma_algebra M" + assumes "finite M" + shows "sigma_algebra \ M" proof - fix A :: "nat \ 'a set" assume "range A \ sets M" - then have "(\i. A i) = (\s\sets M \ range A. s)" + fix A :: "nat \ 'a set" assume "range A \ M" + then have "(\i. A i) = (\s\M \ range A. s)" by auto - also have "(\s\sets M \ range A. s) \ sets M" - using `finite (sets M)` by auto - finally show "(\i. A i) \ sets M" . + also have "(\s\M \ range A. s) \ M" + using `finite M` by auto + finally show "(\i. A i) \ M" . qed lemma countable_UN_eq: fixes A :: "'i::countable \ 'a set" - shows "(range A \ sets M \ (\i. A i) \ sets M) \ - (range (A \ from_nat) \ sets M \ (\i. (A \ from_nat) i) \ sets M)" + shows "(range A \ M \ (\i. A i) \ M) \ + (range (A \ from_nat) \ M \ (\i. (A \ from_nat) i) \ M)" proof - let ?A' = "A \ from_nat" have *: "(\i. ?A' i) = (\i. A i)" (is "?l = ?r") @@ -240,60 +236,57 @@ lemma (in sigma_algebra) countable_UN[intro]: fixes A :: "'i::countable \ 'a set" - assumes "A`X \ sets M" - shows "(\x\X. A x) \ sets M" + assumes "A`X \ M" + shows "(\x\X. A x) \ M" proof - let ?A = "\i. if i \ X then A i else {}" - from assms have "range ?A \ sets M" by auto + from assms have "range ?A \ M" by auto with countable_nat_UN[of "?A \ from_nat"] countable_UN_eq[of ?A M] - have "(\x. ?A x) \ sets M" by auto + have "(\x. ?A x) \ M" by auto moreover have "(\x. ?A x) = (\x\X. A x)" by (auto split: split_if_asm) ultimately show ?thesis by simp qed lemma (in sigma_algebra) countable_INT [intro]: fixes A :: "'i::countable \ 'a set" - assumes A: "A`X \ sets M" "X \ {}" - shows "(\i\X. A i) \ sets M" + assumes A: "A`X \ M" "X \ {}" + shows "(\i\X. A i) \ M" proof - - from A have "\i\X. A i \ sets M" by fast - hence "space M - (\i\X. space M - A i) \ sets M" by blast + from A have "\i\X. A i \ M" by fast + hence "\ - (\i\X. \ - A i) \ M" by blast moreover - have "(\i\X. A i) = space M - (\i\X. space M - A i)" using space_closed A + have "(\i\X. A i) = \ - (\i\X. \ - A i)" using space_closed A by blast ultimately show ?thesis by metis qed -lemma ring_of_sets_Pow: - "ring_of_sets \ space = sp, sets = Pow sp, \ = X \" +lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)" by default auto -lemma algebra_Pow: - "algebra \ space = sp, sets = Pow sp, \ = X \" +lemma algebra_Pow: "algebra sp (Pow sp)" by default auto -lemma sigma_algebra_Pow: - "sigma_algebra \ space = sp, sets = Pow sp, \ = X \" +lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)" by default auto lemma sigma_algebra_iff: - "sigma_algebra M \ - algebra M \ (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" + "sigma_algebra \ M \ + algebra \ M \ (\A. range A \ M \ (\i::nat. A i) \ M)" by (simp add: sigma_algebra_def sigma_algebra_axioms_def) lemma (in sigma_algebra) sets_Collect_countable_All: - assumes "\i. {x\space M. P i x} \ sets M" - shows "{x\space M. \i::'i::countable. P i x} \ sets M" + assumes "\i. {x\\. P i x} \ M" + shows "{x\\. \i::'i::countable. P i x} \ M" proof - - have "{x\space M. \i::'i::countable. P i x} = (\i. {x\space M. P i x})" by auto + have "{x\\. \i::'i::countable. P i x} = (\i. {x\\. P i x})" by auto with assms show ?thesis by auto qed lemma (in sigma_algebra) sets_Collect_countable_Ex: - assumes "\i. {x\space M. P i x} \ sets M" - shows "{x\space M. \i::'i::countable. P i x} \ sets M" + assumes "\i. {x\\. P i x} \ M" + shows "{x\\. \i::'i::countable. P i x} \ M" proof - - have "{x\space M. \i::'i::countable. P i x} = (\i. {x\space M. P i x})" by auto + have "{x\\. \i::'i::countable. P i x} = (\i. {x\\. P i x})" by auto with assms show ?thesis by auto qed @@ -301,9 +294,19 @@ sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All +lemma (in sigma_algebra) sets_Collect_countable_Ball: + assumes "\i. {x\\. P i x} \ M" + shows "{x\\. \i::'i::countable\X. P i x} \ M" + unfolding Ball_def by (intro sets_Collect assms) + +lemma (in sigma_algebra) sets_Collect_countable_Bex: + assumes "\i. {x\\. P i x} \ M" + shows "{x\\. \i::'i::countable\X. P i x} \ M" + unfolding Bex_def by (intro sets_Collect assms) + lemma sigma_algebra_single_set: assumes "X \ S" - shows "sigma_algebra \ space = S, sets = { {}, X, S - X, S }\" + shows "sigma_algebra S { {}, X, S - X, S }" using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \ S`]] by simp subsection {* Binary Unions *} @@ -321,37 +324,34 @@ by (simp add: INF_def range_binary_eq) lemma sigma_algebra_iff2: - "sigma_algebra M \ - sets M \ Pow (space M) \ - {} \ sets M \ (\s \ sets M. space M - s \ sets M) \ - (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" + "sigma_algebra \ M \ + M \ Pow \ \ + {} \ M \ (\s \ M. \ - s \ M) \ + (\A. range A \ M \ (\i::nat. A i) \ M)" by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def algebra_iff_Un Un_range_binary) subsection {* Initial Sigma Algebra *} text {*Sigma algebras can naturally be created as the closure of any set of - sets with regard to the properties just postulated. *} + M with regard to the properties just postulated. *} inductive_set sigma_sets :: "'a set \ 'a set set \ 'a set set" for sp :: "'a set" and A :: "'a set set" where - Basic: "a \ A \ a \ sigma_sets sp A" + Basic[intro, simp]: "a \ A \ a \ sigma_sets sp A" | Empty: "{} \ sigma_sets sp A" | Compl: "a \ sigma_sets sp A \ sp - a \ sigma_sets sp A" | 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), \ = more M \" - lemma (in sigma_algebra) sigma_sets_subset: - assumes a: "a \ sets M" - shows "sigma_sets (space M) a \ sets M" + assumes a: "a \ M" + shows "sigma_sets \ a \ M" proof fix x - assume "x \ sigma_sets (space M) a" - from this show "x \ sets M" + assume "x \ sigma_sets \ a" + from this show "x \ M" by (induct rule: sigma_sets.induct, auto) (metis a subsetD) qed @@ -359,35 +359,28 @@ by (erule sigma_sets.induct, auto) lemma sigma_algebra_sigma_sets: - "a \ Pow (space M) \ sets M = sigma_sets (space M) a \ sigma_algebra M" + "a \ Pow \ \ sigma_algebra \ (sigma_sets \ a)" by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl) lemma sigma_sets_least_sigma_algebra: assumes "A \ Pow S" - shows "sigma_sets S A = \{B. A \ B \ sigma_algebra \space = S, sets = B\}" + shows "sigma_sets S A = \{B. A \ B \ sigma_algebra S B}" proof safe - fix B X assume "A \ B" and sa: "sigma_algebra \ space = S, sets = B \" + fix B X assume "A \ B" and sa: "sigma_algebra S B" and X: "X \ sigma_sets S A" from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \ B`] X show "X \ B" by auto next - fix X assume "X \ \{B. A \ B \ sigma_algebra \space = S, sets = B\}" - then have [intro!]: "\B. A \ B \ sigma_algebra \space = S, sets = B\ \ X \ B" + fix X assume "X \ \{B. A \ B \ sigma_algebra S B}" + then have [intro!]: "\B. A \ B \ sigma_algebra S B \ X \ B" by simp - have "A \ sigma_sets S A" using assms - by (auto intro!: sigma_sets.Basic) - moreover have "sigma_algebra \space = S, sets = sigma_sets S A\" + have "A \ sigma_sets S A" using assms by auto + moreover have "sigma_algebra S (sigma_sets S A)" using assms by (intro sigma_algebra_sigma_sets[of A]) auto ultimately show "X \ sigma_sets S A" by auto qed -lemma sets_sigma: "sets (sigma M) = sigma_sets (space M) (sets M)" - unfolding sigma_def by simp - -lemma space_sigma [simp]: "space (sigma M) = space M" - by (simp add: sigma_def) - lemma sigma_sets_top: "sp \ sigma_sets sp A" by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty) @@ -421,7 +414,7 @@ shows "(\i\S. a i) \ sigma_sets sp A" proof - from ai have "\i. (if i\S then a i else sp) \ sigma_sets sp A" - by (simp add: sigma_sets.intros sigma_sets_top) + by (simp add: sigma_sets.intros(2-) sigma_sets_top) hence "(\i. (if i\S then a i else sp)) \ sigma_sets sp A" by (rule sigma_sets_Inter [OF Asb]) also have "(\i. (if i\S then a i else sp)) = (\i\S. a i)" @@ -430,12 +423,12 @@ qed lemma (in sigma_algebra) sigma_sets_eq: - "sigma_sets (space M) (sets M) = sets M" + "sigma_sets \ M = M" proof - show "sets M \ sigma_sets (space M) (sets M)" + show "M \ sigma_sets \ M" by (metis Set.subsetI sigma_sets.Basic) next - show "sigma_sets (space M) (sets M) \ sets M" + show "sigma_sets \ M \ M" by (metis sigma_sets_subset subset_refl) qed @@ -446,52 +439,42 @@ proof (intro set_eqI iffI) fix a assume "a \ sigma_sets M A" from this A show "a \ sigma_sets M B" - by induct (auto intro!: sigma_sets.intros del: sigma_sets.Basic) + by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic) next fix b assume "b \ sigma_sets M B" from this B show "b \ sigma_sets M A" - by induct (auto intro!: sigma_sets.intros del: sigma_sets.Basic) + by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic) qed -lemma sigma_algebra_sigma: - "sets M \ Pow (space M) \ sigma_algebra (sigma M)" - apply (rule sigma_algebra_sigma_sets) - apply (auto simp add: sigma_def) - done - -lemma (in sigma_algebra) sigma_subset: - "sets N \ sets M \ space N = space M \ sets (sigma N) \ (sets M)" - by (simp add: sigma_def sigma_sets_subset) - 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) + by induct (insert `A \ B`, auto intro: sigma_sets.intros(2-)) qed lemma (in sigma_algebra) restriction_in_sets: fixes A :: "nat \ 'a set" - assumes "S \ sets M" - and *: "range A \ (\A. S \ A) ` sets M" (is "_ \ ?r") - shows "range A \ sets M" "(\i. A i) \ (\A. S \ A) ` sets M" + assumes "S \ M" + and *: "range A \ (\A. S \ A) ` M" (is "_ \ ?r") + shows "range A \ M" "(\i. A i) \ (\A. S \ A) ` M" proof - { fix i have "A i \ ?r" using * by auto - hence "\B. A i = B \ S \ B \ sets M" by auto - hence "A i \ S" "A i \ sets M" using `S \ sets M` by auto } - thus "range A \ sets M" "(\i. A i) \ (\A. S \ A) ` sets M" + hence "\B. A i = B \ S \ B \ M" by auto + hence "A i \ S" "A i \ M" using `S \ M` by auto } + thus "range A \ M" "(\i. A i) \ (\A. S \ A) ` M" by (auto intro!: image_eqI[of _ _ "(\i. A i)"]) qed lemma (in sigma_algebra) restricted_sigma_algebra: - assumes "S \ sets M" - shows "sigma_algebra (restricted_space S)" + assumes "S \ M" + shows "sigma_algebra S (restricted_space S)" unfolding sigma_algebra_def sigma_algebra_axioms_def proof safe - show "algebra (restricted_space S)" using restricted_algebra[OF assms] . + show "algebra S (restricted_space S)" using restricted_algebra[OF assms] . next - fix A :: "nat \ 'a set" assume "range A \ sets (restricted_space S)" + fix A :: "nat \ 'a set" assume "range A \ restricted_space S" from restriction_in_sets[OF assms this[simplified]] - show "(\i. A i) \ sets (restricted_space S)" by simp + show "(\i. A i) \ restricted_space S" by simp qed lemma sigma_sets_Int: @@ -510,7 +493,7 @@ then show ?case by (auto intro!: sigma_sets.Union simp add: UN_extend_simps simp del: UN_simps) - qed (auto intro!: sigma_sets.intros) + qed (auto intro!: sigma_sets.intros(2-)) then show "x \ sigma_sets A (op \ A ` st)" using `A \ sp` by (simp add: Int_absorb2) next @@ -529,93 +512,75 @@ then show ?case by (auto intro!: bexI[of _ "(\x. f x)"] sigma_sets.Union simp add: image_iff) - qed (auto intro!: sigma_sets.intros) + qed (auto intro!: sigma_sets.intros(2-)) qed -lemma sigma_sets_single[simp]: "sigma_sets {X} {{X}} = {{}, {X}}" +lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}" proof (intro set_eqI iffI) - fix x assume "x \ sigma_sets {X} {{X}}" - from sigma_sets_into_sp[OF _ this] - show "x \ {{}, {X}}" by auto + fix a assume "a \ sigma_sets A {}" then show "a \ {{}, A}" + by induct blast+ +qed (auto intro: sigma_sets.Empty sigma_sets_top) + +lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}" +proof (intro set_eqI iffI) + fix x assume "x \ sigma_sets A {A}" + then show "x \ {{}, A}" + by induct blast+ next - fix x assume "x \ {{}, {X}}" - then show "x \ sigma_sets {X} {{X}}" + fix x assume "x \ {{}, A}" + then show "x \ sigma_sets A {A}" by (auto intro: sigma_sets.Empty sigma_sets_top) qed -lemma (in sigma_algebra) sets_sigma_subset: - assumes "space N = space M" - assumes "sets N \ sets M" - shows "sets (sigma N) \ sets M" - by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms) - -lemma in_sigma[intro, simp]: "A \ sets M \ A \ sets (sigma M)" - unfolding sigma_def by (auto intro!: sigma_sets.Basic) - -lemma (in sigma_algebra) sigma_eq[simp]: "sigma M = M" - unfolding sigma_def sigma_sets_eq by simp - -lemma sigma_sigma_eq: - assumes "sets M \ Pow (space M)" - shows "sigma (sigma M) = sigma M" - using sigma_algebra.sigma_eq[OF sigma_algebra_sigma, OF assms] . - lemma sigma_sets_sigma_sets_eq: "M \ Pow S \ sigma_sets S (sigma_sets S M) = sigma_sets S M" - using sigma_sigma_eq[of "\ space = S, sets = M \"] - by (simp add: sigma_def) + by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto lemma sigma_sets_singleton: assumes "X \ S" shows "sigma_sets S { X } = { {}, X, S - X, S }" proof - - interpret sigma_algebra "\ space = S, sets = { {}, X, S - X, S }\" + interpret sigma_algebra S "{ {}, X, S - X, S }" by (rule sigma_algebra_single_set) fact have "sigma_sets S { X } \ sigma_sets S { {}, X, S - X, S }" by (rule sigma_sets_subseteq) simp moreover have "\ = { {}, X, S - X, S }" - using sigma_eq unfolding sigma_def by simp + using sigma_sets_eq by simp moreover { fix A assume "A \ { {}, X, S - X, S }" then have "A \ sigma_sets S { X }" - by (auto intro: sigma_sets.intros sigma_sets_top) } + by (auto intro: sigma_sets.intros(2-) sigma_sets_top) } ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }" by (intro antisym) auto - with sigma_eq show ?thesis - unfolding sigma_def by simp + with sigma_sets_eq show ?thesis by simp qed lemma restricted_sigma: - assumes S: "S \ sets (sigma M)" and M: "sets M \ Pow (space M)" - shows "algebra.restricted_space (sigma M) S = sigma (algebra.restricted_space M S)" + assumes S: "S \ sigma_sets \ M" and M: "M \ Pow \" + shows "algebra.restricted_space (sigma_sets \ M) S = + sigma_sets S (algebra.restricted_space M S)" proof - from S sigma_sets_into_sp[OF M] - have "S \ sigma_sets (space M) (sets M)" "S \ space M" - by (auto simp: sigma_def) + have "S \ sigma_sets \ M" "S \ \" by auto from sigma_sets_Int[OF this] - show ?thesis - by (simp add: sigma_def) + show ?thesis by simp qed lemma sigma_sets_vimage_commute: - assumes X: "X \ space M \ space M'" - shows "{X -` A \ space M |A. A \ sets (sigma M')} - = sigma_sets (space M) {X -` A \ space M |A. A \ sets M'}" (is "?L = ?R") + assumes X: "X \ \ \ \'" + shows "{X -` A \ \ |A. A \ sigma_sets \' M'} + = sigma_sets \ {X -` A \ \ |A. A \ M'}" (is "?L = ?R") proof show "?L \ ?R" proof clarify - fix A assume "A \ sets (sigma M')" - then have "A \ sigma_sets (space M') (sets M')" by (simp add: sets_sigma) - then show "X -` A \ space M \ ?R" + fix A assume "A \ sigma_sets \' M'" + then show "X -` A \ \ \ ?R" proof induct - case (Basic B) then show ?case - by (auto intro!: sigma_sets.Basic) - next case Empty then show ?case by (auto intro!: sigma_sets.Empty) next case (Compl B) - have [simp]: "X -` (space M' - B) \ space M = space M - (X -` B \ space M)" + have [simp]: "X -` (\' - B) \ \ = \ - (X -` B \ \)" by (auto simp add: funcset_mem [OF X]) with Compl show ?case by (auto intro!: sigma_sets.Compl) @@ -624,194 +589,34 @@ then show ?case by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps intro!: sigma_sets.Union) - qed + qed auto qed show "?R \ ?L" proof clarify fix A assume "A \ ?R" - then show "\B. A = X -` B \ space M \ B \ sets (sigma M')" + then show "\B. A = X -` B \ \ \ B \ sigma_sets \' M'" proof induct case (Basic B) then show ?case by auto next case Empty then show ?case - by (auto simp: sets_sigma intro!: sigma_sets.Empty exI[of _ "{}"]) + by (auto intro!: sigma_sets.Empty exI[of _ "{}"]) next case (Compl B) - then obtain A where A: "B = X -` A \ space M" "A \ sets (sigma M')" by auto - then have [simp]: "space M - B = X -` (space M' - A) \ space M" + then obtain A where A: "B = X -` A \ \" "A \ sigma_sets \' M'" by auto + then have [simp]: "\ - B = X -` (\' - A) \ \" by (auto simp add: funcset_mem [OF X]) with A(2) show ?case - by (auto simp: sets_sigma intro: sigma_sets.Compl) + by (auto intro: sigma_sets.Compl) next case (Union F) - then have "\i. \B. F i = X -` B \ space M \ B \ sets (sigma M')" by auto + then have "\i. \B. F i = X -` B \ \ \ B \ sigma_sets \' M'" by auto from choice[OF this] guess A .. note A = this with A show ?case - by (auto simp: sets_sigma vimage_UN[symmetric] intro: sigma_sets.Union) + by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union) qed qed qed -section {* Measurable functions *} - -definition - "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" - -lemma (in sigma_algebra) measurable_sigma: - assumes B: "sets N \ Pow (space N)" - and f: "f \ space M -> space N" - and ba: "\y. y \ sets N \ (f -` y) \ space M \ sets M" - shows "f \ measurable M (sigma N)" -proof - - have "sigma_sets (space N) (sets N) \ {y . (f -` y) \ space M \ sets M & y \ space N}" - proof clarify - fix x - assume "x \ sigma_sets (space N) (sets N)" - thus "f -` x \ space M \ sets M \ x \ space N" - proof induct - case (Basic a) - thus ?case - by (auto simp add: ba) (metis B subsetD PowD) - next - case Empty - thus ?case - by auto - next - case (Compl a) - have [simp]: "f -` space N \ space M = space M" - by (auto simp add: funcset_mem [OF f]) - thus ?case - by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl) - next - case (Union a) - thus ?case - by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast - qed - qed - thus ?thesis - by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f) - (auto simp add: sigma_def) -qed - -lemma measurable_cong: - assumes "\ w. w \ space M \ f w = g w" - shows "f \ measurable M M' \ g \ measurable M M'" - unfolding measurable_def using assms - by (simp cong: vimage_inter_cong Pi_cong) - -lemma measurable_space: - "f \ measurable M A \ x \ space M \ f x \ space A" - unfolding measurable_def by auto - -lemma measurable_sets: - "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" - unfolding measurable_def by auto - -lemma (in sigma_algebra) measurable_subset: - "(\S. S \ sets A \ S \ space A) \ measurable M A \ measurable M (sigma A)" - by (auto intro: measurable_sigma measurable_sets measurable_space) - -lemma measurable_eqI: - "\ space m1 = space m1' ; space m2 = space m2' ; - sets m1 = sets m1' ; sets m2 = sets m2' \ - \ measurable m1 m2 = measurable m1' m2'" - by (simp add: measurable_def sigma_algebra_iff2) - -lemma (in sigma_algebra) measurable_const[intro, simp]: - assumes "c \ space M'" - shows "(\x. c) \ measurable M M'" - using assms by (auto simp add: measurable_def) - -lemma (in sigma_algebra) measurable_If: - assumes measure: "f \ measurable M M'" "g \ measurable M M'" - assumes P: "{x\space M. P x} \ sets M" - shows "(\x. if P x then f x else g x) \ measurable M M'" - unfolding measurable_def -proof safe - fix x assume "x \ space M" - thus "(if P x then f x else g x) \ space M'" - using measure unfolding measurable_def by auto -next - fix A assume "A \ sets M'" - hence *: "(\x. if P x then f x else g x) -` A \ space M = - ((f -` A \ space M) \ {x\space M. P x}) \ - ((g -` A \ space M) \ (space M - {x\space M. P x}))" - using measure unfolding measurable_def by (auto split: split_if_asm) - show "(\x. if P x then f x else g x) -` A \ space M \ sets M" - using `A \ sets M'` measure P unfolding * measurable_def - by (auto intro!: Un) -qed - -lemma (in sigma_algebra) measurable_If_set: - assumes measure: "f \ measurable M M'" "g \ measurable M M'" - assumes P: "A \ sets M" - shows "(\x. if x \ A then f x else g x) \ measurable M M'" -proof (rule measurable_If[OF measure]) - have "{x \ space M. x \ A} = A" using `A \ sets M` sets_into_space by auto - thus "{x \ space M. x \ A} \ sets M" using `A \ sets M` by auto -qed - -lemma (in ring_of_sets) measurable_ident[intro, simp]: "id \ measurable M M" - by (auto simp add: measurable_def) - -lemma measurable_comp[intro]: - fixes f :: "'a \ 'b" and g :: "'b \ 'c" - shows "f \ measurable a b \ g \ measurable b c \ (g o f) \ measurable a c" - apply (auto simp add: measurable_def vimage_compose) - apply (subgoal_tac "f -` g -` y \ space a = f -` (g -` y \ space b) \ space a") - apply force+ - done - -lemma measurable_strong: - fixes f :: "'a \ 'b" and g :: "'b \ 'c" - assumes f: "f \ measurable a b" and g: "g \ (space b -> space c)" - and a: "sigma_algebra a" and b: "sigma_algebra b" and c: "sigma_algebra c" - and t: "f ` (space a) \ t" - and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" - shows "(g o f) \ measurable a c" -proof - - have fab: "f \ (space a -> space b)" - and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f - by (auto simp add: measurable_def) - have eq: "f -` g -` y \ space a = f -` (g -` y \ t) \ space a" using t - by force - show ?thesis - apply (auto simp add: measurable_def vimage_compose a c) - apply (metis funcset_mem fab g) - apply (subst eq, metis ba cb) - done -qed - -lemma measurable_mono1: - "a \ b \ sigma_algebra \space = X, sets = b\ - \ measurable \space = X, sets = a\ c \ measurable \space = X, sets = b\ c" - by (auto simp add: measurable_def) - -lemma measurable_up_sigma: - "measurable A M \ measurable (sigma A) M" - unfolding measurable_def - by (auto simp: sigma_def intro: sigma_sets.Basic) - -lemma (in sigma_algebra) measurable_range_reduce: - "\ f \ measurable M \space = s, sets = Pow s\ ; s \ {} \ - \ f \ measurable M \space = s \ (f ` space M), sets = Pow (s \ (f ` space M))\" - by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast - -lemma (in sigma_algebra) measurable_Pow_to_Pow: - "(sets M = Pow (space M)) \ f \ measurable M \space = UNIV, sets = Pow UNIV\" - by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def) - -lemma (in sigma_algebra) measurable_Pow_to_Pow_image: - "sets M = Pow (space M) - \ f \ measurable M \space = f ` space M, sets = Pow (f ` space M)\" - by (simp add: measurable_def sigma_algebra_Pow) intro_locales - -lemma (in sigma_algebra) measurable_iff_sigma: - assumes "sets E \ Pow (space E)" and "f \ space M \ space E" - shows "f \ measurable M (sigma E) \ (\A\sets E. f -` A \ space M \ sets M)" - using measurable_sigma[OF assms] - by (fastforce simp: measurable_def sets_sigma intro: sigma_sets.intros) - section "Disjoint families" definition @@ -906,8 +711,8 @@ lemma (in ring_of_sets) UNION_in_sets: fixes A:: "nat \ 'a set" - assumes A: "range A \ sets M " - shows "(\i\{0.. sets M" + assumes A: "range A \ M" + shows "(\i\{0.. M" proof (induct n) case 0 show ?case by simp next @@ -917,16 +722,16 @@ qed lemma (in ring_of_sets) range_disjointed_sets: - assumes A: "range A \ sets M " - shows "range (disjointed A) \ sets M" + assumes A: "range A \ M" + shows "range (disjointed A) \ M" proof (auto simp add: disjointed_def) fix n - show "A n - (\i\{0.. sets M" using UNION_in_sets + show "A n - (\i\{0.. M" using UNION_in_sets by (metis A Diff UNIV_I image_subset_iff) qed lemma (in algebra) range_disjointed_sets': - "range A \ sets M \ range (disjointed A) \ sets M" + "range A \ M \ range (disjointed A) \ M" using range_disjointed_sets . lemma disjointed_0[simp]: "disjointed A 0 = A 0" @@ -942,81 +747,518 @@ by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost) lemma sigma_algebra_disjoint_iff: - "sigma_algebra M \ - algebra M & - (\A. range A \ sets M \ disjoint_family A \ - (\i::nat. A i) \ sets M)" + "sigma_algebra \ M \ algebra \ M \ + (\A. range A \ M \ disjoint_family A \ (\i::nat. A i) \ M)" proof (auto simp add: sigma_algebra_iff) fix A :: "nat \ 'a set" - assume M: "algebra M" - and A: "range A \ sets M" - and UnA: "\A. range A \ sets M \ - disjoint_family A \ (\i::nat. A i) \ sets M" - hence "range (disjointed A) \ sets M \ + assume M: "algebra \ M" + and A: "range A \ M" + and UnA: "\A. range A \ M \ disjoint_family A \ (\i::nat. A i) \ M" + hence "range (disjointed A) \ M \ disjoint_family (disjointed A) \ - (\i. disjointed A i) \ sets M" by blast - hence "(\i. disjointed A i) \ sets M" - by (simp add: algebra.range_disjointed_sets' M A disjoint_family_disjointed) - thus "(\i::nat. A i) \ sets M" by (simp add: UN_disjointed_eq) + (\i. disjointed A i) \ M" by blast + hence "(\i. disjointed A i) \ M" + by (simp add: algebra.range_disjointed_sets'[of \] M A disjoint_family_disjointed) + thus "(\i::nat. A i) \ M" by (simp add: UN_disjointed_eq) +qed + +section {* Measure type *} + +definition positive :: "'a set set \ ('a set \ ereal) \ bool" where + "positive M \ \ \ {} = 0 \ (\A\M. 0 \ \ A)" + +definition countably_additive :: "'a set set \ ('a set \ ereal) \ bool" where + "countably_additive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ + (\i. f (A i)) = f (\i. A i))" + +definition measure_space :: "'a set \ 'a set set \ ('a set \ ereal) \ bool" where + "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" + +typedef (open) 'a measure = "{(\::'a set, A, \). (\a\-A. \ a = 0) \ measure_space \ A \ }" +proof + have "sigma_algebra UNIV {{}, UNIV}" + proof qed auto + then show "(UNIV, {{}, UNIV}, \A. 0) \ {(\, A, \). (\a\-A. \ a = 0) \ measure_space \ A \} " + by (auto simp: measure_space_def positive_def countably_additive_def) +qed + +definition space :: "'a measure \ 'a set" where + "space M = fst (Rep_measure M)" + +definition sets :: "'a measure \ 'a set set" where + "sets M = fst (snd (Rep_measure M))" + +definition emeasure :: "'a measure \ 'a set \ ereal" where + "emeasure M = snd (snd (Rep_measure M))" + +definition measure :: "'a measure \ 'a set \ real" where + "measure M A = real (emeasure M A)" + +declare [[coercion sets]] + +declare [[coercion measure]] + +declare [[coercion emeasure]] + +lemma measure_space: "measure_space (space M) (sets M) (emeasure M)" + by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse) + +interpretation sigma_algebra "space M" "sets M" for M :: "'a measure" + using measure_space[of M] by (auto simp: measure_space_def) + +definition measure_of :: "'a set \ 'a set set \ ('a set \ ereal) \ 'a measure" where + "measure_of \ A \ = Abs_measure (\, sigma_sets \ A, + \a. if a \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ a else 0)" + +abbreviation "sigma \ A \ measure_of \ A (\x. 0)" + +lemma measure_space_0: "A \ Pow \ \ measure_space \ (sigma_sets \ A) (\x. 0)" + unfolding measure_space_def + by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def) + +lemma (in ring_of_sets) positive_cong_eq: + "(\a. a \ M \ \' a = \ a) \ positive M \' = positive M \" + by (auto simp add: positive_def) + +lemma (in sigma_algebra) countably_additive_eq: + "(\a. a \ M \ \' a = \ a) \ countably_additive M \' = countably_additive M \" + unfolding countably_additive_def + by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq) + +lemma measure_space_eq: + assumes closed: "A \ Pow \" and eq: "\a. a \ sigma_sets \ A \ \ a = \' a" + shows "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" +proof - + interpret sigma_algebra \ "sigma_sets \ A" using closed by (rule sigma_algebra_sigma_sets) + from positive_cong_eq[OF eq, of "\i. i"] countably_additive_eq[OF eq, of "\i. i"] show ?thesis + by (auto simp: measure_space_def) +qed + +lemma measure_of_eq: + assumes closed: "A \ Pow \" and eq: "(\a. a \ sigma_sets \ A \ \ a = \' a)" + shows "measure_of \ A \ = measure_of \ A \'" +proof - + have "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" + using assms by (rule measure_space_eq) + with eq show ?thesis + by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure]) +qed + +lemma + assumes A: "A \ Pow \" + shows sets_measure_of[simp]: "sets (measure_of \ A \) = sigma_sets \ A" (is ?sets) + and space_measure_of[simp]: "space (measure_of \ A \) = \" (is ?space) +proof - + have "?sets \ ?space" + proof cases + assume "measure_space \ (sigma_sets \ A) \" + moreover have "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) + (\a. if a \ sigma_sets \ A then \ a else 0)" + using A by (rule measure_space_eq) auto + ultimately show "?sets \ ?space" + by (auto simp: Abs_measure_inverse measure_of_def sets_def space_def) + next + assume "\ measure_space \ (sigma_sets \ A) \" + with A show "?sets \ ?space" + by (auto simp: Abs_measure_inverse measure_of_def sets_def space_def measure_space_0) + qed + then show ?sets ?space by auto +qed + +lemma (in sigma_algebra) sets_measure_of_eq[simp]: + "sets (measure_of \ M \) = M" + using space_closed by (auto intro!: sigma_sets_eq) + +lemma (in sigma_algebra) space_measure_of_eq[simp]: + "space (measure_of \ M \) = \" + using space_closed by (auto intro!: sigma_sets_eq) + +lemma measure_of_subset: + "M \ Pow \ \ M' \ M \ sets (measure_of \ M' \) \ sets (measure_of \ M \')" + by (auto intro!: sigma_sets_subseteq) + +lemma in_extended_measure[intro, simp]: "M \ Pow \ \ A \ M \ A \ sets (measure_of \ M \)" + by auto + +section {* Constructing simple @{typ "'a measure"} *} + +lemma emeasure_measure_of: + assumes M: "M = measure_of \ A \" + assumes ms: "A \ Pow \" "positive (sets M) \" "countably_additive (sets M) \" + assumes X: "X \ sets M" + shows "emeasure M X = \ X" +proof - + interpret sigma_algebra \ "sigma_sets \ A" by (rule sigma_algebra_sigma_sets) fact + have "measure_space \ (sigma_sets \ A) \" + using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets) + moreover have "measure_space \ (sigma_sets \ A) (\a. if a \ sigma_sets \ A then \ a else 0) + = measure_space \ (sigma_sets \ A) \" + using ms(1) by (rule measure_space_eq) auto + moreover have "X \ sigma_sets \ A" + using X M ms by simp + ultimately show ?thesis + unfolding emeasure_def measure_of_def M + by (subst Abs_measure_inverse) (simp_all add: sigma_sets_eq) +qed + +lemma emeasure_measure_of_sigma: + assumes ms: "sigma_algebra \ M" "positive M \" "countably_additive M \" + assumes A: "A \ M" + shows "emeasure (measure_of \ M \) A = \ A" +proof - + interpret sigma_algebra \ M by fact + have "measure_space \ (sigma_sets \ M) \" + using ms sigma_sets_eq by (simp add: measure_space_def) + moreover have "measure_space \ (sigma_sets \ M) (\a. if a \ sigma_sets \ M then \ a else 0) + = measure_space \ (sigma_sets \ M) \" + using space_closed by (rule measure_space_eq) auto + ultimately show ?thesis using A + unfolding emeasure_def measure_of_def + by (subst Abs_measure_inverse) (simp_all add: sigma_sets_eq) +qed + +lemma measure_cases[cases type: measure]: + obtains (measure) \ A \ where "x = Abs_measure (\, A, \)" "\a\-A. \ a = 0" "measure_space \ A \" + by atomize_elim (cases x, auto) + +lemma sets_eq_imp_space_eq: + "sets M = sets M' \ space M = space M'" + using top[of M] top[of M'] space_closed[of M] space_closed[of M'] + by blast + +lemma emeasure_notin_sets: "A \ sets M \ emeasure M A = 0" + by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) + +lemma measure_notin_sets: "A \ sets M \ measure M A = 0" + by (simp add: measure_def emeasure_notin_sets) + +lemma measure_eqI: + fixes M N :: "'a measure" + assumes "sets M = sets N" and eq: "\A. A \ sets M \ emeasure M A = emeasure N A" + shows "M = N" +proof (cases M N rule: measure_cases[case_product measure_cases]) + case (measure_measure \ A \ \' A' \') + interpret M: sigma_algebra \ A using measure_measure by (auto simp: measure_space_def) + interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) + have "A = sets M" "A' = sets N" + using measure_measure by (simp_all add: sets_def Abs_measure_inverse) + with `sets M = sets N` have "A = A'" by simp + moreover with M.top N.top M.space_closed N.space_closed have "\ = \'" by auto + moreover { fix B have "\ B = \' B" + proof cases + assume "B \ A" + with eq `A = sets M` have "emeasure M B = emeasure N B" by simp + with measure_measure show "\ B = \' B" + by (simp add: emeasure_def Abs_measure_inverse) + next + assume "B \ A" + with `A = sets M` `A' = sets N` `A = A'` have "B \ sets M" "B \ sets N" + by auto + then have "emeasure M B = 0" "emeasure N B = 0" + by (simp_all add: emeasure_notin_sets) + with measure_measure show "\ B = \' B" + by (simp add: emeasure_def Abs_measure_inverse) + qed } + then have "\ = \'" by auto + ultimately show "M = N" + by (simp add: measure_measure) qed +lemma emeasure_sigma: "A \ Pow \ \ emeasure (sigma \ A) = (\_. 0)" + using measure_space_0[of A \] + by (simp add: measure_of_def emeasure_def Abs_measure_inverse) + +lemma sigma_eqI: + assumes [simp]: "M \ Pow \" "N \ Pow \" "sigma_sets \ M = sigma_sets \ N" + shows "sigma \ M = sigma \ N" + by (rule measure_eqI) (simp_all add: emeasure_sigma) + +section {* Measurable functions *} + +definition measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" where + "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" + +lemma measurable_space: + "f \ measurable M A \ x \ space M \ f x \ space A" + unfolding measurable_def by auto + +lemma measurable_sets: + "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" + unfolding measurable_def by auto + +lemma measurable_sigma_sets: + assumes B: "sets N = sigma_sets \ A" "A \ Pow \" + and f: "f \ space M \ \" + and ba: "\y. y \ A \ (f -` y) \ space M \ sets M" + shows "f \ measurable M N" +proof - + interpret A: sigma_algebra \ "sigma_sets \ A" using B(2) by (rule sigma_algebra_sigma_sets) + from B top[of N] A.top space_closed[of N] A.space_closed have \: "\ = space N" by force + + { fix X assume "X \ sigma_sets \ A" + then have "f -` X \ space M \ sets M \ X \ \" + proof induct + case (Basic a) then show ?case + by (auto simp add: ba) (metis B(2) subsetD PowD) + next + case (Compl a) + have [simp]: "f -` \ \ space M = space M" + by (auto simp add: funcset_mem [OF f]) + then show ?case + by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl) + next + case (Union a) + then show ?case + by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast + qed auto } + with f show ?thesis + by (auto simp add: measurable_def B \) +qed + +lemma measurable_measure_of: + assumes B: "N \ Pow \" + and f: "f \ space M \ \" + and ba: "\y. y \ N \ (f -` y) \ space M \ sets M" + shows "f \ measurable M (measure_of \ N \)" +proof - + have "sets (measure_of \ N \) = sigma_sets \ N" + using B by (rule sets_measure_of) + from this assms show ?thesis by (rule measurable_sigma_sets) +qed + +lemma measurable_iff_measure_of: + assumes "N \ Pow \" "f \ space M \ \" + shows "f \ measurable M (measure_of \ N \) \ (\A\N. f -` A \ space M \ sets M)" + by (metis assms in_extended_measure measurable_measure_of assms measurable_sets) + +lemma measurable_cong: + assumes "\ w. w \ space M \ f w = g w" + shows "f \ measurable M M' \ g \ measurable M M'" + unfolding measurable_def using assms + by (simp cong: vimage_inter_cong Pi_cong) + +lemma measurable_eqI: + "\ space m1 = space m1' ; space m2 = space m2' ; + sets m1 = sets m1' ; sets m2 = sets m2' \ + \ measurable m1 m2 = measurable m1' m2'" + by (simp add: measurable_def sigma_algebra_iff2) + +lemma measurable_const[intro, simp]: + "c \ space M' \ (\x. c) \ measurable M M'" + by (auto simp add: measurable_def) + +lemma measurable_If: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "{x\space M. P x} \ sets M" + shows "(\x. if P x then f x else g x) \ measurable M M'" + unfolding measurable_def +proof safe + fix x assume "x \ space M" + thus "(if P x then f x else g x) \ space M'" + using measure unfolding measurable_def by auto +next + fix A assume "A \ sets M'" + hence *: "(\x. if P x then f x else g x) -` A \ space M = + ((f -` A \ space M) \ {x\space M. P x}) \ + ((g -` A \ space M) \ (space M - {x\space M. P x}))" + using measure unfolding measurable_def by (auto split: split_if_asm) + show "(\x. if P x then f x else g x) -` A \ space M \ sets M" + using `A \ sets M'` measure P unfolding * measurable_def + by (auto intro!: Un) +qed + +lemma measurable_If_set: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "A \ sets M" + shows "(\x. if x \ A then f x else g x) \ measurable M M'" +proof (rule measurable_If[OF measure]) + have "{x \ space M. x \ A} = A" using `A \ sets M` sets_into_space by auto + thus "{x \ space M. x \ A} \ sets M" using `A \ sets M` by auto +qed + +lemma measurable_ident[intro, simp]: "id \ measurable M M" + by (auto simp add: measurable_def) + +lemma measurable_comp[intro]: + fixes f :: "'a \ 'b" and g :: "'b \ 'c" + shows "f \ measurable a b \ g \ measurable b c \ (g o f) \ measurable a c" + apply (auto simp add: measurable_def vimage_compose) + apply (subgoal_tac "f -` g -` y \ space a = f -` (g -` y \ space b) \ space a") + apply force+ + done + +lemma measurable_Least: + assumes meas: "\i::nat. {x\space M. P i x} \ M" + shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" +proof - + { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" + proof cases + assume i: "(LEAST j. False) = i" + have "(\x. LEAST j. P j x) -` {i} \ space M = + {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" + by (simp add: set_eq_iff, safe) + (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) + with meas show ?thesis + by (auto intro!: Int) + next + assume i: "(LEAST j. False) \ i" + then have "(\x. LEAST j. P j x) -` {i} \ space M = + {x\space M. P i x} \ (space M - (\jspace M. P j x}))" + proof (simp add: set_eq_iff, safe) + fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" + have "\j. P j x" + by (rule ccontr) (insert neq, auto) + then show "P (LEAST j. P j x) x" by (rule LeastI_ex) + qed (auto dest: Least_le intro!: Least_equality) + with meas show ?thesis + by auto + qed } + then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" + by (intro countable_UN) auto + moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = + (\x. LEAST j. P j x) -` A \ space M" by auto + ultimately show ?thesis by auto +qed + +lemma measurable_strong: + fixes f :: "'a \ 'b" and g :: "'b \ 'c" + assumes f: "f \ measurable a b" and g: "g \ space b \ space c" + and t: "f ` (space a) \ t" + and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" + shows "(g o f) \ measurable a c" +proof - + have fab: "f \ (space a -> space b)" + and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f + by (auto simp add: measurable_def) + have eq: "\y. f -` g -` y \ space a = f -` (g -` y \ t) \ space a" using t + by force + show ?thesis + apply (auto simp add: measurable_def vimage_compose) + apply (metis funcset_mem fab g) + apply (subst eq, metis ba cb) + done +qed + +lemma measurable_mono1: + "M' \ Pow \ \ M \ M' \ + measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" + using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) + +subsection {* Extend measure *} + +definition "extend_measure \ I G \ = + (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) + then measure_of \ (G`I) (SOME \'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') + else measure_of \ (G`I) (\_. 0))" + +lemma space_extend_measure: "G ` I \ Pow \ \ space (extend_measure \ I G \) = \" + unfolding extend_measure_def by simp + +lemma sets_extend_measure: "G ` I \ Pow \ \ sets (extend_measure \ I G \) = sigma_sets \ (G`I)" + unfolding extend_measure_def by simp + +lemma emeasure_extend_measure: + assumes M: "M = extend_measure \ I G \" + and eq: "\i. i \ I \ \' (G i) = \ i" + and ms: "G ` I \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" + and "i \ I" + shows "emeasure M (G i) = \ i" +proof cases + assume *: "(\i\I. \ i = 0)" + with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" + by (simp add: extend_measure_def) + from measure_space_0[OF ms(1)] ms `i\I` + have "emeasure M (G i) = 0" + by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) + with `i\I` * show ?thesis + by simp +next + def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" + assume "\ (\i\I. \ i = 0)" + moreover + have "measure_space (space M) (sets M) \'" + using ms unfolding measure_space_def by auto default + with ms eq have "\\'. P \'" + unfolding P_def + by (intro exI[of _ \']) (auto simp add: M space_extend_measure sets_extend_measure) + ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" + by (simp add: M extend_measure_def P_def[symmetric]) + + from `\\'. P \'` have P: "P (Eps P)" by (rule someI_ex) + show "emeasure M (G i) = \ i" + proof (subst emeasure_measure_of[OF M_eq]) + have sets_M: "sets M = sigma_sets \ (G`I)" + using M_eq ms by (auto simp: sets_extend_measure) + then show "G i \ sets M" using `i \ I` by auto + show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" + using P `i\I` by (auto simp add: sets_M measure_space_def P_def) + qed fact +qed + +lemma emeasure_extend_measure_Pair: + assumes M: "M = extend_measure \ {(i, j). I i j} (\(i, j). G i j) (\(i, j). \ i j)" + and eq: "\i j. I i j \ \' (G i j) = \ i j" + and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" + and "I i j" + shows "emeasure M (G i j) = \ i j" + using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j` + by (auto simp: subset_eq) + subsection {* Sigma algebra generated by function preimages *} -definition (in sigma_algebra) - "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)" - by (simp add: vimage_algebra_def image_iff) +definition + "vimage_algebra M S f = sigma S ((\A. f -` A \ S) ` sets M)" -lemma (in sigma_algebra) space_vimage_algebra[simp]: - "space (vimage_algebra S f) = S" - by (simp add: vimage_algebra_def) - -lemma (in sigma_algebra) sigma_algebra_preimages: +lemma sigma_algebra_preimages: fixes f :: "'x \ 'a" - assumes "f \ A \ space M" - shows "sigma_algebra \ space = A, sets = (\M. f -` M \ A) ` sets M \" - (is "sigma_algebra \ space = _, sets = ?F ` sets M \") + assumes "f \ S \ space M" + shows "sigma_algebra S ((\A. f -` A \ S) ` sets M)" + (is "sigma_algebra _ (?F ` sets M)") proof (simp add: sigma_algebra_iff2, safe) show "{} \ ?F ` sets M" by blast next - fix S assume "S \ sets M" - moreover have "A - ?F S = ?F (space M - S)" + fix A assume "A \ sets M" + moreover have "S - ?F A = ?F (space M - A)" using assms by auto - ultimately show "A - ?F S \ ?F ` sets M" + ultimately show "S - ?F A \ ?F ` sets M" by blast next - fix S :: "nat \ 'x set" assume *: "range S \ ?F ` sets M" - have "\i. \b. b \ sets M \ S i = ?F b" + fix A :: "nat \ 'x set" assume *: "range A \ ?F ` M" + have "\i. \b. b \ M \ A i = ?F b" proof safe fix i - have "S i \ ?F ` sets M" using * by auto - then show "\b. b \ sets M \ S i = ?F b" by auto + have "A i \ ?F ` M" using * by auto + then show "\b. b \ M \ A i = ?F b" by auto qed - from choice[OF this] obtain b where b: "range b \ sets M" "\i. S i = ?F (b i)" + from choice[OF this] obtain b where b: "range b \ M" "\i. A i = ?F (b i)" by auto - then have "(\i. S i) = ?F (\i. b i)" by auto - then show "(\i. S i) \ ?F ` sets M" using b(1) by blast + then have "(\i. A i) = ?F (\i. b i)" by auto + then show "(\i. A i) \ ?F ` M" using b(1) by blast qed -lemma (in sigma_algebra) sigma_algebra_vimage: +lemma sets_vimage_algebra[simp]: + "f \ S \ space M \ sets (vimage_algebra M S f) = (\A. f -` A \ S) ` sets M" + using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M] + by (simp add: vimage_algebra_def) + +lemma space_vimage_algebra[simp]: + "f \ S \ space M \ space (vimage_algebra M S f) = S" + using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M] + by (simp add: vimage_algebra_def) + +lemma in_vimage_algebra[simp]: + "f \ S \ space M \ A \ sets (vimage_algebra M S f) \ (\B\sets M. A = f -` B \ S)" + by (simp add: image_iff) + +lemma measurable_vimage_algebra: fixes S :: "'c set" assumes "f \ S \ space M" - shows "sigma_algebra (vimage_algebra S f)" -proof - - from sigma_algebra_preimages[OF assms] - show ?thesis unfolding vimage_algebra_def by (auto simp: sigma_algebra_iff2) -qed + shows "f \ measurable (vimage_algebra M S f) M" + unfolding measurable_def using assms by force -lemma (in sigma_algebra) measurable_vimage_algebra: - fixes S :: "'c set" assumes "f \ S \ space M" - shows "f \ measurable (vimage_algebra S f) M" - unfolding measurable_def using assms by force - -lemma (in sigma_algebra) measurable_vimage: +lemma measurable_vimage: fixes g :: "'a \ 'c" and f :: "'d \ 'a" assumes "g \ measurable M M2" "f \ S \ space M" - shows "(\x. g (f x)) \ measurable (vimage_algebra S f) M2" + shows "(\x. g (f x)) \ measurable (vimage_algebra M S f) M2" proof - note measurable_vimage_algebra[OF assms(2)] from measurable_comp[OF this assms(1)] @@ -1033,7 +1275,7 @@ proof induct case (Basic X) then obtain X' where "X = ?F X'" "X' \ A" by auto - then show ?case by (auto intro!: sigma_sets.Basic) + then show ?case by auto next case Empty then show ?case by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty) @@ -1060,8 +1302,6 @@ then obtain X' where "X' \ sigma_sets S A" "X = ?F X'" by auto then show "X \ sigma_sets S' (?F ` A)" proof (induct arbitrary: X) - case (Basic X') then show ?case by (auto intro: sigma_sets.Basic) - next case Empty then show ?case by (auto intro: sigma_sets.Empty) next case (Compl X') @@ -1078,23 +1318,7 @@ also have "(\i. f -` F i \ S') = X" using assms Union by auto finally show ?case . - qed -qed - -section {* Conditional space *} - -definition (in algebra) - "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" - -lemma (in algebra) space_conditional_space: - assumes "A \ sets M" shows "space (conditional_space X A) = X`A" -proof - - interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra) - show ?thesis unfolding conditional_space_def r.image_space_def - by simp + qed auto qed subsection {* A Two-Element Series *} @@ -1113,80 +1337,64 @@ section {* Closed CDI *} -definition - closed_cdi where - "closed_cdi M \ - sets M \ Pow (space M) & - (\s \ sets M. space M - s \ sets M) & - (\A. (range A \ sets M) & (A 0 = {}) & (\n. A n \ A (Suc n)) \ - (\i. A i) \ sets M) & - (\A. (range A \ sets M) & disjoint_family A \ (\i::nat. A i) \ sets M)" +definition closed_cdi where + "closed_cdi \ M \ + M \ Pow \ & + (\s \ M. \ - s \ M) & + (\A. (range A \ M) & (A 0 = {}) & (\n. A n \ A (Suc n)) \ + (\i. A i) \ M) & + (\A. (range A \ M) & disjoint_family A \ (\i::nat. A i) \ M)" inductive_set - smallest_ccdi_sets :: "('a, 'b) algebra_scheme \ 'a set set" - for M + smallest_ccdi_sets :: "'a set \ 'a set set \ 'a set set" + for \ M where Basic [intro]: - "a \ sets M \ a \ smallest_ccdi_sets M" + "a \ M \ a \ smallest_ccdi_sets \ M" | Compl [intro]: - "a \ smallest_ccdi_sets M \ space M - a \ smallest_ccdi_sets M" + "a \ smallest_ccdi_sets \ M \ \ - a \ smallest_ccdi_sets \ M" | Inc: - "range A \ Pow(smallest_ccdi_sets M) \ A 0 = {} \ (\n. A n \ A (Suc n)) - \ (\i. A i) \ smallest_ccdi_sets M" + "range A \ Pow(smallest_ccdi_sets \ M) \ A 0 = {} \ (\n. A n \ A (Suc n)) + \ (\i. A i) \ smallest_ccdi_sets \ M" | Disj: - "range A \ Pow(smallest_ccdi_sets M) \ disjoint_family A - \ (\i::nat. A i) \ smallest_ccdi_sets M" - + "range A \ Pow(smallest_ccdi_sets \ M) \ disjoint_family A + \ (\i::nat. A i) \ smallest_ccdi_sets \ M" -definition - smallest_closed_cdi where - "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)" +lemma (in subset_class) smallest_closed_cdi1: "M \ smallest_ccdi_sets \ M" + by auto -lemma space_smallest_closed_cdi [simp]: - "space (smallest_closed_cdi M) = space M" - by (simp add: smallest_closed_cdi_def) - -lemma (in algebra) smallest_closed_cdi1: "sets M \ sets (smallest_closed_cdi M)" - by (auto simp add: smallest_closed_cdi_def) - -lemma (in algebra) smallest_ccdi_sets: - "smallest_ccdi_sets M \ Pow (space M)" +lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \ M \ Pow \" apply (rule subsetI) apply (erule smallest_ccdi_sets.induct) apply (auto intro: range_subsetD dest: sets_into_space) done -lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)" - apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets) +lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \ (smallest_ccdi_sets \ M)" + apply (auto simp add: closed_cdi_def smallest_ccdi_sets) apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) + done -lemma (in algebra) smallest_closed_cdi3: - "sets (smallest_closed_cdi M) \ Pow (space M)" - by (simp add: smallest_closed_cdi_def smallest_ccdi_sets) - -lemma closed_cdi_subset: "closed_cdi M \ sets M \ Pow (space M)" +lemma closed_cdi_subset: "closed_cdi \ M \ M \ Pow \" by (simp add: closed_cdi_def) -lemma closed_cdi_Compl: "closed_cdi M \ s \ sets M \ space M - s \ sets M" +lemma closed_cdi_Compl: "closed_cdi \ M \ s \ M \ \ - s \ M" by (simp add: closed_cdi_def) lemma closed_cdi_Inc: - "closed_cdi M \ range A \ sets M \ A 0 = {} \ (!!n. A n \ A (Suc n)) \ - (\i. A i) \ sets M" + "closed_cdi \ M \ range A \ M \ A 0 = {} \ (!!n. A n \ A (Suc n)) \ (\i. A i) \ M" by (simp add: closed_cdi_def) lemma closed_cdi_Disj: - "closed_cdi M \ range A \ sets M \ disjoint_family A \ (\i::nat. A i) \ sets M" + "closed_cdi \ M \ range A \ M \ disjoint_family A \ (\i::nat. A i) \ M" by (simp add: closed_cdi_def) lemma closed_cdi_Un: - assumes cdi: "closed_cdi M" and empty: "{} \ sets M" - and A: "A \ sets M" and B: "B \ sets M" + assumes cdi: "closed_cdi \ M" and empty: "{} \ M" + and A: "A \ M" and B: "B \ M" and disj: "A \ B = {}" - shows "A \ B \ sets M" + shows "A \ B \ M" proof - - have ra: "range (binaryset A B) \ sets M" + have ra: "range (binaryset A B) \ M" by (simp add: range_binaryset_eq empty A B) have di: "disjoint_family (binaryset A B)" using disj by (simp add: disjoint_family_on_def binaryset_def Int_commute) @@ -1196,11 +1404,11 @@ qed lemma (in algebra) smallest_ccdi_sets_Un: - assumes A: "A \ smallest_ccdi_sets M" and B: "B \ smallest_ccdi_sets M" + assumes A: "A \ smallest_ccdi_sets \ M" and B: "B \ smallest_ccdi_sets \ M" and disj: "A \ B = {}" - shows "A \ B \ smallest_ccdi_sets M" + shows "A \ B \ smallest_ccdi_sets \ M" proof - - have ra: "range (binaryset A B) \ Pow (smallest_ccdi_sets M)" + have ra: "range (binaryset A B) \ Pow (smallest_ccdi_sets \ M)" by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic) have di: "disjoint_family (binaryset A B)" using disj by (simp add: disjoint_family_on_def binaryset_def Int_commute) @@ -1210,33 +1418,32 @@ qed lemma (in algebra) smallest_ccdi_sets_Int1: - assumes a: "a \ sets M" - shows "b \ smallest_ccdi_sets M \ a \ b \ smallest_ccdi_sets M" + assumes a: "a \ M" + shows "b \ smallest_ccdi_sets \ M \ a \ b \ smallest_ccdi_sets \ M" proof (induct rule: smallest_ccdi_sets.induct) case (Basic x) thus ?case by (metis a Int smallest_ccdi_sets.Basic) next case (Compl x) - have "a \ (space M - x) = space M - ((space M - a) \ (a \ x))" + have "a \ (\ - x) = \ - ((\ - a) \ (a \ x))" by blast - also have "... \ smallest_ccdi_sets M" + also have "... \ smallest_ccdi_sets \ M" by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2 - Diff_disjoint Int_Diff Int_empty_right Un_commute - smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl - smallest_ccdi_sets_Un) + Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un + smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl) finally show ?case . next case (Inc A) have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" by blast - have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets M)" using Inc + have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets \ M)" using Inc by blast moreover have "(\i. a \ A i) 0 = {}" by (simp add: Inc) moreover have "!!n. (\i. a \ A i) n \ (\i. a \ A i) (Suc n)" using Inc by blast - ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets M" + ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets \ M" by (rule smallest_ccdi_sets.Inc) show ?case by (metis 1 2) @@ -1244,11 +1451,11 @@ case (Disj A) have 1: "(\i. (\i. a \ A i) i) = a \ (\i. A i)" by blast - have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets M)" using Disj + have "range (\i. a \ A i) \ Pow(smallest_ccdi_sets \ M)" using Disj by blast moreover have "disjoint_family (\i. a \ A i)" using Disj by (auto simp add: disjoint_family_on_def) - ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets M" + ultimately have 2: "(\i. (\i. a \ A i) i) \ smallest_ccdi_sets \ M" by (rule smallest_ccdi_sets.Disj) show ?case by (metis 1 2) @@ -1256,17 +1463,17 @@ lemma (in algebra) smallest_ccdi_sets_Int: - assumes b: "b \ smallest_ccdi_sets M" - shows "a \ smallest_ccdi_sets M \ a \ b \ smallest_ccdi_sets M" + assumes b: "b \ smallest_ccdi_sets \ M" + shows "a \ smallest_ccdi_sets \ M \ a \ b \ smallest_ccdi_sets \ M" proof (induct rule: smallest_ccdi_sets.induct) case (Basic x) thus ?case by (metis b smallest_ccdi_sets_Int1) next case (Compl x) - have "(space M - x) \ b = space M - (x \ b \ (space M - b))" + have "(\ - x) \ b = \ - (x \ b \ (\ - b))" by blast - also have "... \ smallest_ccdi_sets M" + also have "... \ smallest_ccdi_sets \ M" by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b smallest_ccdi_sets.Compl smallest_ccdi_sets_Un) finally show ?case . @@ -1274,13 +1481,13 @@ case (Inc A) have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" by blast - have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets M)" using Inc + have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets \ M)" using Inc by blast moreover have "(\i. A i \ b) 0 = {}" by (simp add: Inc) moreover have "!!n. (\i. A i \ b) n \ (\i. A i \ b) (Suc n)" using Inc by blast - ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets M" + ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets \ M" by (rule smallest_ccdi_sets.Inc) show ?case by (metis 1 2) @@ -1288,39 +1495,34 @@ case (Disj A) have 1: "(\i. (\i. A i \ b) i) = (\i. A i) \ b" by blast - have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets M)" using Disj + have "range (\i. A i \ b) \ Pow(smallest_ccdi_sets \ M)" using Disj by blast moreover have "disjoint_family (\i. A i \ b)" using Disj by (auto simp add: disjoint_family_on_def) - ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets M" + ultimately have 2: "(\i. (\i. A i \ b) i) \ smallest_ccdi_sets \ M" by (rule smallest_ccdi_sets.Disj) show ?case by (metis 1 2) qed -lemma (in algebra) sets_smallest_closed_cdi_Int: - "a \ sets (smallest_closed_cdi M) \ b \ sets (smallest_closed_cdi M) - \ a \ b \ sets (smallest_closed_cdi M)" - by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def) - lemma (in algebra) sigma_property_disjoint_lemma: - assumes sbC: "sets M \ C" - and ccdi: "closed_cdi (|space = space M, sets = C|)" - shows "sigma_sets (space M) (sets M) \ C" + assumes sbC: "M \ C" + and ccdi: "closed_cdi \ C" + shows "sigma_sets \ M \ C" proof - - have "smallest_ccdi_sets M \ {B . sets M \ B \ sigma_algebra (|space = space M, sets = B|)}" + have "smallest_ccdi_sets \ M \ {B . M \ B \ sigma_algebra \ B}" apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int smallest_ccdi_sets_Int) apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets) apply (blast intro: smallest_ccdi_sets.Disj) done - hence "sigma_sets (space M) (sets M) \ smallest_ccdi_sets M" + hence "sigma_sets (\) (M) \ smallest_ccdi_sets \ M" by clarsimp - (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto) + (drule sigma_algebra.sigma_sets_subset [where a="M"], auto) also have "... \ C" proof fix x - assume x: "x \ smallest_ccdi_sets M" + assume x: "x \ smallest_ccdi_sets \ M" thus "x \ C" proof (induct rule: smallest_ccdi_sets.induct) case (Basic x) @@ -1344,21 +1546,21 @@ qed lemma (in algebra) sigma_property_disjoint: - assumes sbC: "sets M \ C" - and compl: "!!s. s \ C \ sigma_sets (space M) (sets M) \ space M - s \ C" - and inc: "!!A. range A \ C \ sigma_sets (space M) (sets M) + assumes sbC: "M \ C" + and compl: "!!s. s \ C \ sigma_sets (\) (M) \ \ - s \ C" + and inc: "!!A. range A \ C \ sigma_sets (\) (M) \ A 0 = {} \ (!!n. A n \ A (Suc n)) \ (\i. A i) \ C" - and disj: "!!A. range A \ C \ sigma_sets (space M) (sets M) + and disj: "!!A. range A \ C \ sigma_sets (\) (M) \ disjoint_family A \ (\i::nat. A i) \ C" - shows "sigma_sets (space M) (sets M) \ C" + shows "sigma_sets (\) (M) \ C" proof - - have "sigma_sets (space M) (sets M) \ C \ sigma_sets (space M) (sets M)" + have "sigma_sets (\) (M) \ C \ sigma_sets (\) (M)" proof (rule sigma_property_disjoint_lemma) - show "sets M \ C \ sigma_sets (space M) (sets M)" + show "M \ C \ sigma_sets (\) (M)" by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic) next - show "closed_cdi \space = space M, sets = C \ sigma_sets (space M) (sets M)\" + show "closed_cdi \ (C \ sigma_sets (\) (M))" by (simp add: closed_cdi_def compl inc disj) (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed IntE sigma_sets.Compl range_subsetD sigma_sets.Union) @@ -1370,97 +1572,97 @@ section {* Dynkin systems *} locale dynkin_system = subset_class + - assumes space: "space M \ sets M" - and compl[intro!]: "\A. A \ sets M \ space M - A \ sets M" - and UN[intro!]: "\A. disjoint_family A \ range A \ sets M - \ (\i::nat. A i) \ sets M" + assumes space: "\ \ M" + and compl[intro!]: "\A. A \ M \ \ - A \ M" + and UN[intro!]: "\A. disjoint_family A \ range A \ M + \ (\i::nat. A i) \ M" -lemma (in dynkin_system) empty[intro, simp]: "{} \ sets M" - using space compl[of "space M"] by simp +lemma (in dynkin_system) empty[intro, simp]: "{} \ M" + using space compl[of "\"] by simp lemma (in dynkin_system) diff: - assumes sets: "D \ sets M" "E \ sets M" and "D \ E" - shows "E - D \ sets M" + assumes sets: "D \ M" "E \ M" and "D \ E" + shows "E - D \ M" proof - - let ?f = "\x. if x = 0 then D else if x = Suc 0 then space M - E else {}" - have "range ?f = {D, space M - E, {}}" + let ?f = "\x. if x = 0 then D else if x = Suc 0 then \ - E else {}" + have "range ?f = {D, \ - E, {}}" by (auto simp: image_iff) - moreover have "D \ (space M - E) = (\i. ?f i)" + moreover have "D \ (\ - E) = (\i. ?f i)" by (auto simp: image_iff split: split_if_asm) moreover then have "disjoint_family ?f" unfolding disjoint_family_on_def - using `D \ sets M`[THEN sets_into_space] `D \ E` by auto - ultimately have "space M - (D \ (space M - E)) \ sets M" + using `D \ M`[THEN sets_into_space] `D \ E` by auto + ultimately have "\ - (D \ (\ - E)) \ M" using sets by auto - also have "space M - (D \ (space M - E)) = E - D" + also have "\ - (D \ (\ - E)) = E - D" using assms sets_into_space by auto finally show ?thesis . qed lemma dynkin_systemI: - assumes "\ A. A \ sets M \ A \ space M" "space M \ sets M" - assumes "\ A. A \ sets M \ space M - A \ sets M" - assumes "\ A. disjoint_family A \ range A \ sets M - \ (\i::nat. A i) \ sets M" - shows "dynkin_system M" + assumes "\ A. A \ M \ A \ \" "\ \ M" + assumes "\ A. A \ M \ \ - A \ M" + assumes "\ A. disjoint_family A \ range A \ M + \ (\i::nat. A i) \ M" + shows "dynkin_system \ M" using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def) lemma dynkin_systemI': - assumes 1: "\ A. A \ sets M \ A \ space M" - assumes empty: "{} \ sets M" - assumes Diff: "\ A. A \ sets M \ space M - A \ sets M" - assumes 2: "\ A. disjoint_family A \ range A \ sets M - \ (\i::nat. A i) \ sets M" - shows "dynkin_system M" + assumes 1: "\ A. A \ M \ A \ \" + assumes empty: "{} \ M" + assumes Diff: "\ A. A \ M \ \ - A \ M" + assumes 2: "\ A. disjoint_family A \ range A \ M + \ (\i::nat. A i) \ M" + shows "dynkin_system \ M" proof - - from Diff[OF empty] have "space M \ sets M" by auto + from Diff[OF empty] have "\ \ M" by auto from 1 this Diff 2 show ?thesis by (intro dynkin_systemI) auto qed lemma dynkin_system_trivial: - shows "dynkin_system \ space = A, sets = Pow A \" + shows "dynkin_system A (Pow A)" by (rule dynkin_systemI) auto lemma sigma_algebra_imp_dynkin_system: - assumes "sigma_algebra M" shows "dynkin_system M" + assumes "sigma_algebra \ M" shows "dynkin_system \ M" proof - - interpret sigma_algebra M by fact + interpret sigma_algebra \ M by fact show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI) qed subsection "Intersection stable algebras" -definition "Int_stable M \ (\ a \ sets M. \ b \ sets M. a \ b \ sets M)" +definition "Int_stable M \ (\ a \ M. \ b \ M. a \ b \ M)" lemma (in algebra) Int_stable: "Int_stable M" unfolding Int_stable_def by auto lemma Int_stableI: - "(\a b. a \ A \ b \ A \ a \ b \ A) \ Int_stable \ space = \, sets = A \" + "(\a b. a \ A \ b \ A \ a \ b \ A) \ Int_stable A" unfolding Int_stable_def by auto lemma Int_stableD: - "Int_stable M \ a \ sets M \ b \ sets M \ a \ b \ sets M" + "Int_stable M \ a \ M \ b \ M \ a \ b \ M" unfolding Int_stable_def by auto lemma (in dynkin_system) sigma_algebra_eq_Int_stable: - "sigma_algebra M \ Int_stable M" + "sigma_algebra \ M \ Int_stable M" proof - assume "sigma_algebra M" then show "Int_stable M" + assume "sigma_algebra \ M" then show "Int_stable M" unfolding sigma_algebra_def using algebra.Int_stable by auto next assume "Int_stable M" - show "sigma_algebra M" + show "sigma_algebra \ M" unfolding sigma_algebra_disjoint_iff algebra_iff_Un proof (intro conjI ballI allI impI) - show "sets M \ Pow (space M)" using sets_into_space by auto + show "M \ Pow (\)" using sets_into_space by auto next - fix A B assume "A \ sets M" "B \ sets M" - then have "A \ B = space M - ((space M - A) \ (space M - B))" - "space M - A \ sets M" "space M - B \ sets M" + fix A B assume "A \ M" "B \ M" + then have "A \ B = \ - ((\ - A) \ (\ - B))" + "\ - A \ M" "\ - B \ M" using sets_into_space by auto - then show "A \ B \ sets M" + then show "A \ B \ M" using `Int_stable M` unfolding Int_stable_def by auto qed auto qed @@ -1468,217 +1670,148 @@ subsection "Smallest Dynkin systems" definition dynkin where - "dynkin M = \ space = space M, - sets = \{D. dynkin_system \ space = space M, sets = D \ \ sets M \ D}, - \ = more M \" + "dynkin \ M = (\{D. dynkin_system \ D \ M \ D})" lemma dynkin_system_dynkin: - assumes "sets M \ Pow (space M)" - shows "dynkin_system (dynkin M)" + assumes "M \ Pow (\)" + shows "dynkin_system \ (dynkin \ M)" proof (rule dynkin_systemI) - fix A assume "A \ sets (dynkin M)" + fix A assume "A \ dynkin \ M" moreover - { fix D assume "A \ D" and d: "dynkin_system \ space = space M, sets = D \" - then have "A \ space M" by (auto simp: dynkin_system_def subset_class_def) } - moreover have "{D. dynkin_system \ space = space M, sets = D\ \ sets M \ D} \ {}" + { fix D assume "A \ D" and d: "dynkin_system \ D" + then have "A \ \" by (auto simp: dynkin_system_def subset_class_def) } + moreover have "{D. dynkin_system \ D \ M \ D} \ {}" using assms dynkin_system_trivial by fastforce - ultimately show "A \ space (dynkin M)" + ultimately show "A \ \" unfolding dynkin_def using assms - by simp (metis dynkin_system_def subset_class_def in_mono) + by auto next - show "space (dynkin M) \ sets (dynkin M)" + show "\ \ dynkin \ M" unfolding dynkin_def using dynkin_system.space by fastforce next - fix A assume "A \ sets (dynkin M)" - then show "space (dynkin M) - A \ sets (dynkin M)" + fix A assume "A \ dynkin \ M" + then show "\ - A \ dynkin \ M" unfolding dynkin_def using dynkin_system.compl by force next fix A :: "nat \ 'a set" - assume A: "disjoint_family A" "range A \ sets (dynkin M)" - show "(\i. A i) \ sets (dynkin M)" unfolding dynkin_def + assume A: "disjoint_family A" "range A \ dynkin \ M" + show "(\i. A i) \ dynkin \ M" unfolding dynkin_def proof (simp, safe) - fix D assume "dynkin_system \space = space M, sets = D\" "sets M \ D" - with A have "(\i. A i) \ sets \space = space M, sets = D\" + fix D assume "dynkin_system \ D" "M \ D" + with A have "(\i. A i) \ D" by (intro dynkin_system.UN) (auto simp: dynkin_def) then show "(\i. A i) \ D" by auto qed qed -lemma dynkin_Basic[intro]: - "A \ sets M \ A \ sets (dynkin M)" - unfolding dynkin_def by auto - -lemma dynkin_space[simp]: - "space (dynkin M) = space M" +lemma dynkin_Basic[intro]: "A \ M \ A \ dynkin \ M" unfolding dynkin_def by auto lemma (in dynkin_system) restricted_dynkin_system: - assumes "D \ sets M" - shows "dynkin_system \ space = space M, - sets = {Q. Q \ space M \ Q \ D \ sets M} \" + assumes "D \ M" + shows "dynkin_system \ {Q. Q \ \ \ Q \ D \ M}" proof (rule dynkin_systemI, simp_all) - have "space M \ D = D" - using `D \ sets M` sets_into_space by auto - then show "space M \ D \ sets M" - using `D \ sets M` by auto + have "\ \ D = D" + using `D \ M` sets_into_space by auto + then show "\ \ D \ M" + using `D \ M` by auto next - fix A assume "A \ space M \ A \ D \ sets M" - moreover have "(space M - A) \ D = (space M - (A \ D)) - (space M - D)" + fix A assume "A \ \ \ A \ D \ M" + moreover have "(\ - A) \ D = (\ - (A \ D)) - (\ - D)" by auto - ultimately show "space M - A \ space M \ (space M - A) \ D \ sets M" - using `D \ sets M` by (auto intro: diff) + ultimately show "\ - A \ \ \ (\ - A) \ D \ M" + using `D \ M` by (auto intro: diff) next fix A :: "nat \ 'a set" - assume "disjoint_family A" "range A \ {Q. Q \ space M \ Q \ D \ sets M}" - then have "\i. A i \ space M" "disjoint_family (\i. A i \ D)" - "range (\i. A i \ D) \ sets M" "(\x. A x) \ D = (\x. A x \ D)" + assume "disjoint_family A" "range A \ {Q. Q \ \ \ Q \ D \ M}" + then have "\i. A i \ \" "disjoint_family (\i. A i \ D)" + "range (\i. A i \ D) \ M" "(\x. A x) \ D = (\x. A x \ D)" by ((fastforce simp: disjoint_family_on_def)+) - then show "(\x. A x) \ space M \ (\x. A x) \ D \ sets M" + then show "(\x. A x) \ \ \ (\x. A x) \ D \ M" by (auto simp del: UN_simps) qed lemma (in dynkin_system) dynkin_subset: - assumes "sets N \ sets M" - assumes "space N = space M" - shows "sets (dynkin N) \ sets M" + assumes "N \ M" + shows "dynkin \ N \ M" proof - - have "dynkin_system M" by default - then have "dynkin_system \space = space N, sets = sets M \" + have "dynkin_system \ M" by default + then have "dynkin_system \ M" using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp - with `sets N \ sets M` show ?thesis by (auto simp add: dynkin_def) + with `N \ M` show ?thesis by (auto simp add: dynkin_def) qed lemma sigma_eq_dynkin: - assumes sets: "sets M \ Pow (space M)" + assumes sets: "M \ Pow \" assumes "Int_stable M" - shows "sigma M = dynkin M" + shows "sigma_sets \ M = dynkin \ M" proof - - have "sets (dynkin M) \ sigma_sets (space M) (sets M)" + have "dynkin \ M \ sigma_sets (\) (M)" using sigma_algebra_imp_dynkin_system - unfolding dynkin_def sigma_def sigma_sets_least_sigma_algebra[OF sets] by auto + unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto moreover - interpret dynkin_system "dynkin M" + interpret dynkin_system \ "dynkin \ M" using dynkin_system_dynkin[OF sets] . - have "sigma_algebra (dynkin M)" + have "sigma_algebra \ (dynkin \ M)" unfolding sigma_algebra_eq_Int_stable Int_stable_def proof (intro ballI) - fix A B assume "A \ sets (dynkin M)" "B \ sets (dynkin M)" - let ?D = "\E. \ space = space M, - sets = {Q. Q \ space M \ Q \ E \ sets (dynkin M)} \" - have "sets M \ sets (?D B)" + fix A B assume "A \ dynkin \ M" "B \ dynkin \ M" + let ?D = "\E. {Q. Q \ \ \ Q \ E \ dynkin \ M}" + have "M \ ?D B" proof - fix E assume "E \ sets M" - then have "sets M \ sets (?D E)" "E \ sets (dynkin M)" + fix E assume "E \ M" + then have "M \ ?D E" "E \ dynkin \ M" using sets_into_space `Int_stable M` by (auto simp: Int_stable_def) - then have "sets (dynkin M) \ sets (?D E)" - using restricted_dynkin_system `E \ sets (dynkin M)` + then have "dynkin \ M \ ?D E" + using restricted_dynkin_system `E \ dynkin \ M` by (intro dynkin_system.dynkin_subset) simp_all - then have "B \ sets (?D E)" - using `B \ sets (dynkin M)` by auto - then have "E \ B \ sets (dynkin M)" + then have "B \ ?D E" + using `B \ dynkin \ M` by auto + then have "E \ B \ dynkin \ M" by (subst Int_commute) simp - then show "E \ sets (?D B)" - using sets `E \ sets M` by auto + then show "E \ ?D B" + using sets `E \ M` by auto qed - then have "sets (dynkin M) \ sets (?D B)" - using restricted_dynkin_system `B \ sets (dynkin M)` + then have "dynkin \ M \ ?D B" + using restricted_dynkin_system `B \ dynkin \ M` by (intro dynkin_system.dynkin_subset) simp_all - then show "A \ B \ sets (dynkin M)" - using `A \ sets (dynkin M)` sets_into_space by auto + then show "A \ B \ dynkin \ M" + using `A \ dynkin \ M` sets_into_space by auto qed - from sigma_algebra.sigma_sets_subset[OF this, of "sets M"] - 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 + from sigma_algebra.sigma_sets_subset[OF this, of "M"] + have "sigma_sets (\) (M) \ dynkin \ M" by auto + ultimately have "sigma_sets (\) (M) = dynkin \ M" by auto then show ?thesis - by (auto intro!: algebra.equality simp: sigma_def dynkin_def) + by (auto simp: dynkin_def) qed lemma (in dynkin_system) dynkin_idem: - "dynkin M = M" + "dynkin \ M = M" proof - - have "sets (dynkin M) = sets M" + have "dynkin \ M = M" proof - show "sets M \ sets (dynkin M)" + show "M \ dynkin \ M" using dynkin_Basic by auto - show "sets (dynkin M) \ sets M" + show "dynkin \ M \ M" by (intro dynkin_subset) auto qed then show ?thesis - by (auto intro!: algebra.equality simp: dynkin_def) + by (auto simp: dynkin_def) qed lemma (in dynkin_system) dynkin_lemma: 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" + and E: "E \ M" "M \ sigma_sets \ E" + shows "sigma_sets \ E = M" proof - - have "sets E \ Pow (space E)" + have "E \ Pow \" using E sets_into_space by force - then have "sigma E = dynkin E" + then have "sigma_sets \ 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(1,2)] by simp + moreover then have "dynkin \ E = M" + using assms dynkin_subset[OF E(1)] by simp ultimately show ?thesis - using assms by (auto intro!: algebra.equality simp: dynkin_def) -qed - -subsection "Sigma algebras on finite sets" - -locale finite_sigma_algebra = sigma_algebra + - assumes finite_space: "finite (space M)" - and sets_eq_Pow[simp]: "sets M = Pow (space M)" - -lemma (in finite_sigma_algebra) sets_image_space_eq_Pow: - "sets (image_space X) = Pow (space (image_space X))" -proof safe - fix x S assume "S \ sets (image_space X)" "x \ S" - then show "x \ space (image_space X)" - using sets_into_space by (auto intro!: imageI simp: image_space_def) -next - fix S assume "S \ space (image_space X)" - then obtain S' where "S = X`S'" "S'\sets M" - by (auto simp: subset_image_iff image_space_def) - then show "S \ sets (image_space X)" - by (auto simp: image_space_def) -qed - -lemma measurable_sigma_sigma: - assumes M: "sets M \ Pow (space M)" and N: "sets N \ Pow (space N)" - shows "f \ measurable M N \ f \ measurable (sigma M) (sigma N)" - using sigma_algebra.measurable_subset[OF sigma_algebra_sigma[OF M], of N] - using measurable_up_sigma[of M N] N by auto - -lemma (in sigma_algebra) measurable_Least: - assumes meas: "\i::nat. {x\space M. P i x} \ sets M" - shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" -proof - - { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" - proof cases - assume i: "(LEAST j. False) = i" - have "(\x. LEAST j. P j x) -` {i} \ space M = - {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" - by (simp add: set_eq_iff, safe) - (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) - with meas show ?thesis - by (auto intro!: Int) - next - assume i: "(LEAST j. False) \ i" - then have "(\x. LEAST j. P j x) -` {i} \ space M = - {x\space M. P i x} \ (space M - (\jspace M. P j x}))" - proof (simp add: set_eq_iff, safe) - fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" - have "\j. P j x" - by (rule ccontr) (insert neq, auto) - then show "P (LEAST j. P j x) x" by (rule LeastI_ex) - qed (auto dest: Least_le intro!: Least_equality) - with meas show ?thesis - by (auto intro!: Int) - qed } - then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" - by (intro countable_UN) auto - moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = - (\x. LEAST j. P j x) -` A \ space M" by auto - ultimately show ?thesis by auto + using assms by (auto simp: dynkin_def) qed end diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Apr 23 12:14:35 2012 +0200 @@ -8,29 +8,6 @@ lemma Ex1_eq: "\! x. P x \ P x \ P y \ x = y" by auto -locale finite_space = - fixes S :: "'a set" - assumes finite[simp]: "finite S" - and not_empty[simp]: "S \ {}" - -definition (in finite_space) "M = \ space = S, sets = Pow S, - measure = \A. ereal (card A / card S) \" - -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 - by (rule finite_measure_spaceI) - (simp_all add: M_def real_of_nat_def) - -sublocale finite_space \ information_space M 2 - by default (simp_all add: M_def one_ereal_def) - -lemma (in finite_space) \'_eq[simp]: "\' A = (if A \ S then card A / card S else 0)" - unfolding \'_def by (auto simp: M_def) - section "Define the state space" text {* @@ -68,7 +45,7 @@ definition "dining_cryptographers = ({None} \ Some ` {0.. {xs :: bool list. length xs = n}" definition "payer dc = fst dc" -definition coin :: "(nat option \ bool list) => nat \ bool" where +definition coin :: "(nat option \ bool list) \ nat \ bool" where "coin dc c = snd dc ! (c mod n)" definition "inversion dc = map (\c. (payer dc = Some c) \ (coin dc c \ coin dc (c + 1))) [0.. prob_space "uniform_count_measure dc_crypto" + by (rule prob_space_uniform_count_measure[OF finite_dc_crypto]) + (insert n_gt_3, auto simp: dc_crypto intro: exI[of _ "replicate n True"]) -sublocale - dining_cryptographers_space \ finite_space "dc_crypto" -proof - show "finite dc_crypto" using finite_dc_crypto . - show "dc_crypto \ {}" - unfolding dc_crypto - using n_gt_3 by (auto intro: exI[of _ "replicate n True"]) -qed +sublocale dining_cryptographers_space \ information_space "uniform_count_measure dc_crypto" 2 + by default auto notation (in dining_cryptographers_space) mutual_information_Pow ("\'( _ ; _ ')") @@ -438,71 +412,64 @@ theorem (in dining_cryptographers_space) "\( inversion ; payer ) = 0" -proof - +proof (rule mutual_information_eq_0_simple) have n: "0 < n" using n_gt_3 by auto - have lists: "{xs. length xs = n} \ {}" using Ex_list_of_length by auto have card_image_inversion: "real (card (inversion ` dc_crypto)) = 2^n / 2" unfolding card_image_inversion using `0 < n` by (cases n) auto - let ?dIP = "distribution (\x. (inversion x, payer x))" - let ?dP = "distribution payer" - let ?dI = "distribution inversion" + show inversion: "simple_distributed (uniform_count_measure dc_crypto) inversion (\x. 2 / 2^n)" + proof (rule simple_distributedI) + show "simple_function (uniform_count_measure dc_crypto) inversion" + using finite_dc_crypto + by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) + fix x assume "x \ inversion ` space (uniform_count_measure dc_crypto)" + moreover have "inversion -` {x} \ dc_crypto = {dc \ dc_crypto. inversion dc = x}" by auto + ultimately show "2 / 2^n = prob (inversion -` {x} \ space (uniform_count_measure dc_crypto))" + using `0 < n` + by (simp add: card_inversion card_dc_crypto finite_dc_crypto + subset_eq space_uniform_count_measure measure_uniform_count_measure) + qed - { have "\(inversion|payer) = - - (\x\inversion`dc_crypto. (\z\Some ` {0..x\inversion`dc_crypto. (\z\Some ` {0.. inversion`dc_crypto" and z: "z \ Some ` {0..x. (inversion x, payer x)) -` {(x, z)} \ dc_crypto = - {dc \ dc_crypto. payer dc = Some (the z) \ inversion dc = x}" - by (auto simp add: payer_def) - moreover from x z obtain i where "z = Some i" and "i < n" by auto - moreover from x have "length x = n" by (auto simp: inversion_def [abs_def] dc_crypto) - ultimately - have "?dIP {(x, z)} = 2 / (real n * 2^n)" using x - by (auto simp add: card_dc_crypto card_payer_and_inversion distribution_def) - moreover - from z have "payer -` {z} \ dc_crypto = {z} \ {xs. length xs = n}" - by (auto simp: dc_crypto payer_def) - hence "card (payer -` {z} \ dc_crypto) = 2^n" - using card_lists_length_eq[where A="UNIV::bool set"] - by (simp add: card_cartesian_product_singleton) - hence "?dP {z} = 1 / real n" - by (simp add: distribution_def card_dc_crypto) - ultimately - show "?dIP {(x,z)} * log 2 (?dIP {(x,z)} / ?dP {z}) = - 2 / (real n * 2^n) * (1 - real n)" - by (simp add: log_divide log_nat_power[of 2]) - qed - also have "... = real n - 1" - using n finite_space - by (simp add: card_image_inversion card_image[OF inj_Some] field_simps real_eq_of_nat[symmetric]) - finally have "\(inversion|payer) = real n - 1" . } - moreover - { have "\(inversion) = - (\x \ inversion`dc_crypto. ?dI {x} * log 2 (?dI {x}))" - unfolding entropy_eq[OF simple_function_finite] by simp - also have "... = - (\x \ inversion`dc_crypto. 2 * (1 - real n) / 2^n)" - unfolding neg_equal_iff_equal - proof (rule setsum_cong[OF refl]) - fix x assume x_inv: "x \ inversion ` dc_crypto" - hence "length x = n" by (auto simp: inversion_def [abs_def] 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 (auto simp: card_inversion[OF x_inv] card_dc_crypto distribution_def) - thus "?dI {x} * log 2 (?dI {x}) = 2 * (1 - real n) / 2^n" - by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide) - qed - also have "... = real n - 1" - by (simp add: card_image_inversion real_of_nat_def[symmetric] field_simps) - finally have "\(inversion) = real n - 1" . - } - ultimately show ?thesis - unfolding mutual_information_eq_entropy_conditional_entropy[OF simple_function_finite simple_function_finite] + show "simple_distributed (uniform_count_measure dc_crypto) payer (\x. 1 / real n)" + proof (rule simple_distributedI) + show "simple_function (uniform_count_measure dc_crypto) payer" + using finite_dc_crypto + by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) + fix z assume "z \ payer ` space (uniform_count_measure dc_crypto)" + then have "payer -` {z} \ dc_crypto = {z} \ {xs. length xs = n}" + by (auto simp: dc_crypto payer_def space_uniform_count_measure) + hence "card (payer -` {z} \ dc_crypto) = 2^n" + using card_lists_length_eq[where A="UNIV::bool set"] + by (simp add: card_cartesian_product_singleton) + then show "1 / real n = prob (payer -` {z} \ space (uniform_count_measure dc_crypto))" + using finite_dc_crypto + by (subst measure_uniform_count_measure) (auto simp add: card_dc_crypto space_uniform_count_measure) + qed + + show "simple_distributed (uniform_count_measure dc_crypto) (\x. (inversion x, payer x)) (\x. 2 / (real n *2^n))" + proof (rule simple_distributedI) + show "simple_function (uniform_count_measure dc_crypto) (\x. (inversion x, payer x))" + using finite_dc_crypto + by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure) + fix x assume "x \ (\x. (inversion x, payer x)) ` space (uniform_count_measure dc_crypto)" + then obtain i xs where x: "x = (inversion (Some i, xs), payer (Some i, xs))" + and "i < n" "length xs = n" + by (simp add: image_iff space_uniform_count_measure dc_crypto Bex_def) blast + then have xs: "inversion (Some i, xs) \ inversion`dc_crypto" and i: "Some i \ Some ` {0.. "inversion (Some i, xs)" + ultimately have ys: "ys \ inversion`dc_crypto" + and "Some i \ Some ` {0..x. (inversion x, payer x)) -` {x} \ space (uniform_count_measure dc_crypto) = + {dc \ dc_crypto. payer dc = Some (the (Some i)) \ inversion dc = ys}" + by (auto simp add: payer_def space_uniform_count_measure) + then show "2 / (real n * 2 ^ n) = prob ((\x. (inversion x, payer x)) -` {x} \ space (uniform_count_measure dc_crypto))" + using `i < n` ys + by (simp add: measure_uniform_count_measure card_payer_and_inversion finite_dc_crypto subset_eq card_dc_crypto) + qed + + show "\x\space (uniform_count_measure dc_crypto). 2 / (real n * 2 ^ n) = 2 / 2 ^ n * (1 / real n) " by simp qed diff -r 64023cf4d148 -r 05663f75964c src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Apr 23 12:23:23 2012 +0100 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Apr 23 12:14:35 2012 +0200 @@ -103,49 +103,14 @@ lessThan_Suc_eq_insert_0 setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric]) qed -lemma ex_ordered_bij_betw_nat_finite: - fixes order :: "nat \ 'a\linorder" - assumes "finite S" - shows "\f. bij_betw f {0.. (\ij j \ order (f i) \ order (f j))" -proof - - from ex_bij_betw_nat_finite [OF `finite S`] guess f .. note f = this - let ?xs = "sort_key order (map f [0 ..< card S])" +declare space_point_measure[simp] - have "?xs <~~> map f [0 ..< card S]" - unfolding multiset_of_eq_perm[symmetric] by (rule multiset_of_sort) - from permutation_Ex_bij [OF this] - obtain g where g: "bij_betw g {0..i. i ?xs ! i = map f [0 ..< card S] ! g i" - by (auto simp: atLeast0LessThan) - - { fix i assume "i < card S" - then have "g i < card S" using g by (auto simp: bij_betw_def) - with map [OF `i < card S`] have "f (g i) = ?xs ! i" by simp } - note this[simp] +declare sets_point_measure[simp] - show ?thesis - proof (intro exI allI conjI impI) - show "bij_betw (f\g) {0.. j" - from sorted_nth_mono[of "map order ?xs" i j] - show "order ((f\g) i) \ order ((f\g) j)" by simp - qed -qed - -definition (in prob_space) - "ordered_variable_partition X = (SOME f. bij_betw f {0.. - (\ij j \ distribution X {f i} \ distribution X {f j}))" - -definition (in prob_space) - "order_distribution X i = real (distribution X {ordered_variable_partition X i})" - -definition (in prob_space) - "guessing_entropy b X = (\i'(_')") where - "\(X) \ guessing_entropy b X" +lemma measure_point_measure: + "finite \ \ A \ \ \ (\x. x \ \ \ 0 \ p x) \ + measure (point_measure \ (\x. ereal (p x))) A = (\i\A. p i)" + unfolding measure_def by (subst emeasure_point_measure_finite) auto locale finite_information = fixes \ :: "'a set" @@ -159,17 +124,14 @@ lemma (in finite_information) positive_p_sum[simp]: "0 \ setsum p X" by (auto intro!: setsum_nonneg) -sublocale finite_information \ finite_measure_space "\ space = \, sets = Pow \, measure = \S. ereal (setsum p S)\" - by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset) +sublocale finite_information \ prob_space "point_measure \ p" + by default (simp add: one_ereal_def emeasure_point_measure_finite) -sublocale finite_information \ finite_prob_space "\ space = \, sets = Pow \, measure = \S. ereal (setsum p S)\" - by default (simp add: one_ereal_def) - -sublocale finite_information \ information_space "\ space = \, sets = Pow \, measure = \S. ereal (setsum p S)\" b +sublocale finite_information \ information_space "point_measure \ p" b by default simp -lemma (in finite_information) \'_eq: "A \ \ \ \' A = setsum p A" - unfolding \'_def by auto +lemma (in finite_information) \'_eq: "A \ \ \ prob A = setsum p A" + by (auto simp: measure_point_measure) locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b for b :: real @@ -247,14 +209,19 @@ abbreviation "p A \ setsum P A" +abbreviation + "\ \ point_measure msgs P" + abbreviation probability ("\

'(_') _") where - "\

(X) x \ distribution X x" + "\

(X) x \ measure \ (X -` x \ msgs)" -abbreviation joint_probability ("\

'(_, _') _") where - "\

(X, Y) x \ joint_distribution X Y x" +abbreviation joint_probability ("\

'(_ ; _') _") where + "\

(X ; Y) x \ \

(\x. (X x, Y x)) x" -abbreviation conditional_probability ("\

'(_|_') _") where - "\

(X|Y) x \ \

(X, Y) x / \

(Y) (snd`x)" +no_notation disj (infixr "|" 30) + +abbreviation conditional_probability ("\

'(_ | _') _") where + "\

(X | Y) x \ (\

(X ; Y) x) / \

(Y) (snd`x)" notation entropy_Pow ("\'( _ ')") @@ -280,8 +247,8 @@ from assms have *: "fst -` {k} \ msgs = {k}\{ms. set ms \ messages \ length ms = n}" unfolding msgs_def by auto - show "\

(fst) {k} = K k" - apply (simp add: \'_eq distribution_def) + show "(\

(fst) {k}) = K k" + apply (simp add: \'_eq) apply (simp add: * P_def) apply (simp add: setsum_cartesian_product') using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\x x. True"] `k \ keys` @@ -297,6 +264,67 @@ unfolding msgs_def fst_image_times if_not_P[OF *] by simp qed +lemma setsum_distribution_cut: + "\

(X) {x} = (\y \ Y`space \. \

(X ; Y) {(x, y)})" + by (subst finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def inj_on_def + intro!: arg_cong[where f=prob]) + +lemma prob_conj_imp1: + "prob ({x. Q x} \ msgs) = 0 \ prob ({x. Pr x \ Q x} \ msgs) = 0" + using finite_measure_mono[of "{x. Pr x \ Q x} \ msgs" "{x. Q x} \ msgs"] + using measure_nonneg[of \ "{x. Pr x \ Q x} \ msgs"] + by (simp add: subset_eq) + +lemma prob_conj_imp2: + "prob ({x. Pr x} \ msgs) = 0 \ prob ({x. Pr x \ Q x} \ msgs) = 0" + using finite_measure_mono[of "{x. Pr x \ Q x} \ msgs" "{x. Pr x} \ msgs"] + using measure_nonneg[of \ "{x. Pr x \ Q x} \ msgs"] + by (simp add: subset_eq) + +lemma simple_function_finite: "simple_function \ f" + by (simp add: simple_function_def) + +lemma entropy_commute: "\(\x. (X x, Y x)) = \(\x. (Y x, X x))" + apply (subst (1 2) entropy_simple_distributed[OF simple_distributedI[OF simple_function_finite refl]]) + unfolding space_point_measure +proof - + have eq: "(\x. (X x, Y x)) ` msgs = (\(x, y). (y, x)) ` (\x. (Y x, X x)) ` msgs" + by auto + show "- (\x\(\x. (X x, Y x)) ` msgs. (\

(X ; Y) {x}) * log b (\

(X ; Y) {x})) = + - (\x\(\x. (Y x, X x)) ` msgs. (\

(Y ; X) {x}) * log b (\

(Y ; X) {x}))" + unfolding eq + apply (subst setsum_reindex) + apply (auto simp: vimage_def inj_on_def intro!: setsum_cong arg_cong[where f="\x. prob x * log b (prob x)"]) + done +qed + +lemma (in -) measure_eq_0I: "A = {} \ measure M A = 0" by simp + +lemma conditional_entropy_eq_ce_with_hypothesis: + "\(X | Y) = -(\y\Y`msgs. (\

(Y) {y}) * (\x\X`msgs. (\

(X ; Y) {(x,y)}) / (\

(Y) {y}) * + log b ((\

(X ; Y) {(x,y)}) / (\

(Y) {y}))))" + apply (subst conditional_entropy_eq[OF + simple_distributedI[OF simple_function_finite refl] + simple_function_finite + simple_distributedI[OF simple_function_finite refl]]) + unfolding space_point_measure +proof - + have "- (\(x, y)\(\x. (X x, Y x)) ` msgs. (\

(X ; Y) {(x, y)}) * log b ((\

(X ; Y) {(x, y)}) / (\

(Y) {y}))) = + - (\x\X`msgs. (\y\Y`msgs. (\

(X ; Y) {(x, y)}) * log b ((\

(X ; Y) {(x, y)}) / (\

(Y) {y}))))" + unfolding setsum_cartesian_product + apply (intro arg_cong[where f=uminus] setsum_mono_zero_left) + apply (auto simp: vimage_def image_iff intro!: measure_eq_0I) + apply metis + done + also have "\ = - (\y\Y`msgs. (\x\X`msgs. (\

(X ; Y) {(x, y)}) * log b ((\

(X ; Y) {(x, y)}) / (\

(Y) {y}))))" + by (subst setsum_commute) rule + also have "\ = -(\y\Y`msgs. (\

(Y) {y}) * (\x\X`msgs. (\

(X ; Y) {(x,y)}) / (\

(Y) {y}) * log b ((\

(X ; Y) {(x,y)}) / (\

(Y) {y}))))" + by (auto simp add: setsum_right_distrib vimage_def intro!: setsum_cong prob_conj_imp1) + finally show "- (\(x, y)\(\x. (X x, Y x)) ` msgs. (\

(X ; Y) {(x, y)}) * log b ((\

(X ; Y) {(x, y)}) / (\

(Y) {y}))) = + -(\y\Y`msgs. (\

(Y) {y}) * (\x\X`msgs. (\

(X ; Y) {(x,y)}) / (\

(Y) {y}) * log b ((\

(X ; Y) {(x,y)}) / (\

(Y) {y}))))" . +qed + lemma ce_OB_eq_ce_t: "\(fst | OB) = \(fst | t\OB)" proof - txt {* Lemma 2 *} @@ -314,22 +342,22 @@ then have **: "\ms. length ms = n \ OB (k, ms) = obs \ (\i(OB, fst) {(obs, k)} / K k = + have "(\

(OB ; fst) {(obs, k)}) / K k = p ({k}\{ms. (k,ms) \ msgs \ OB (k,ms) = obs}) / K k" - apply (simp add: distribution_def \'_eq) by (auto intro!: arg_cong[where f=p]) + apply (simp add: \'_eq) by (auto intro!: arg_cong[where f=p]) also have "\ = (\im\{m\messages. observe k m = obs ! i}. M m)" unfolding P_def using `K k \ 0` `k \ keys` apply (simp add: setsum_cartesian_product' setsum_divide_distrib msgs_def ** cong: conj_cong) apply (subst setprod_setsum_distrib_lists[OF M.finite_space]) .. - finally have "\

(OB, fst) {(obs, k)} / K k = + finally have "(\

(OB ; fst) {(obs, k)}) / K k = (\im\{m\messages. observe k m = obs ! i}. M m)" . } note * = this - have "\

(OB, fst) {(obs, k)} / K k = \

(OB, fst) {(obs', k)} / K k" + have "(\

(OB ; fst) {(obs, k)}) / K k = (\

(OB ; fst) {(obs', k)}) / K k" unfolding *[OF obs] *[OF obs'] using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: setprod_reindex) - then have "\

(OB, fst) {(obs, k)} = \

(OB, fst) {(obs', k)}" + then have "(\

(OB ; fst) {(obs, k)}) = (\

(OB ; fst) {(obs', k)})" using `K k \ 0` by auto } note t_eq_imp = this @@ -339,13 +367,13 @@ (\obs'\?S obs. ((\x. (OB x, fst x)) -` {(obs', k)} \ msgs))" by auto have df: "disjoint_family_on (\obs'. (\x. (OB x, fst x)) -` {(obs', k)} \ msgs) (?S obs)" unfolding disjoint_family_on_def by auto - have "\

(t\OB, fst) {(t obs, k)} = (\obs'\?S obs. \

(OB, fst) {(obs', k)})" - unfolding distribution_def comp_def - using finite_measure_finite_Union[OF _ _ df] - by (force simp add: * intro!: setsum_nonneg) - also have "(\obs'\?S obs. \

(OB, fst) {(obs', k)}) = real (card (?S obs)) * \

(OB, fst) {(obs, k)}" + have "\

(t\OB ; fst) {(t obs, k)} = (\obs'\?S obs. \

(OB ; fst) {(obs', k)})" + unfolding comp_def + using finite_measure_finite_Union[OF _ df] + by (auto simp add: * intro!: setsum_nonneg) + also have "(\obs'\?S obs. \

(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" by (simp add: t_eq_imp[OF `k \ keys` `K k \ 0` obs] real_eq_of_nat) - finally have "\

(t\OB, fst) {(t obs, k)} = real (card (?S obs)) * \

(OB, fst) {(obs, k)}" .} + finally have "\

(t\OB ; fst) {(t obs, k)} = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" .} note P_t_eq_P_OB = this { fix k obs assume "k \ keys" and obs: "obs \ OB`msgs" @@ -359,11 +387,15 @@ then have "real (card ?S) \ 0" by auto have "\

(fst | t\OB) {(k, t obs)} = \

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

(fst) {k} / \

(t\OB) {t obs}" - using distribution_order(7,8)[where X=fst and x=k and Y="t\OB" and y="t obs"] - by (subst joint_distribution_commute) auto - also have "\

(t\OB) {t obs} = (\k'\keys. \

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

(fst) {k'})" - using setsum_distribution_cut(2)[of "t\OB" fst "t obs", symmetric] - by (auto intro!: setsum_cong distribution_order(8)) + using finite_measure_mono[of "{x. fst x = k \ t (OB x) = t obs} \ msgs" "{x. fst x = k} \ msgs"] + using measure_nonneg[of \ "{x. fst x = k \ t (OB x) = t obs} \ msgs"] + by (auto simp add: vimage_def conj_commute subset_eq) + also have "(\

(t\OB) {t obs}) = (\k'\keys. (\

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

(fst) {k'}))" + using finite_measure_mono[of "{x. t (OB x) = t obs \ fst x = k} \ msgs" "{x. fst x = k} \ msgs"] + using measure_nonneg[of \ "{x. fst x = k \ t (OB x) = t obs} \ msgs"] + apply (simp add: setsum_distribution_cut[of "t\OB" "t obs" fst]) + apply (auto intro!: setsum_cong simp: subset_eq vimage_def prob_conj_imp1) + done also have "\

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

(fst) {k} / (\k'\keys. \

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

(fst) {k'}) = \

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

(fst) {k} / (\k'\keys. \

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

(fst) {k'})" using CP_t_K[OF `k\keys` obs] CP_t_K[OF _ obs] `real (card ?S) \ 0` @@ -371,10 +403,10 @@ mult_divide_mult_cancel_left[OF `real (card ?S) \ 0`] cong: setsum_cong) also have "(\k'\keys. \

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

(fst) {k'}) = \

(OB) {obs}" - using setsum_distribution_cut(2)[of OB fst obs, symmetric] - by (auto intro!: setsum_cong distribution_order(8)) + using setsum_distribution_cut[of OB obs fst] + by (auto intro!: setsum_cong simp: prob_conj_imp1 vimage_def) also have "\

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

(fst) {k} / \

(OB) {obs} = \

(fst | OB) {(k, obs)}" - by (subst joint_distribution_commute) (auto intro!: distribution_order(8)) + by (auto simp: vimage_def conj_commute prob_conj_imp2) finally have "\

(fst | t\OB) {(k, t obs)} = \

(fst | OB) {(k, obs)}" . } note CP_T_eq_CP_O = this @@ -396,15 +428,15 @@ have df: "disjoint_family_on (\obs. OB -` {obs} \ msgs) (?S (OB x))" unfolding disjoint_family_on_def by auto have "\

(t\OB) {t (OB x)} = (\obs\?S (OB x). \

(OB) {obs})" - unfolding distribution_def comp_def - using finite_measure_finite_Union[OF _ _ df] + unfolding comp_def + using finite_measure_finite_Union[OF _ df] by (force simp add: * intro!: setsum_nonneg) } note P_t_sum_P_O = this txt {* Lemma 3 *} + txt {* Lemma 3 *} have "\(fst | OB) = -(\obs\OB`msgs. \

(OB) {obs} * ?Ht (t obs))" - unfolding conditional_entropy_eq_ce_with_hypothesis[OF - simple_function_finite simple_function_finite] using * by simp + unfolding conditional_entropy_eq_ce_with_hypothesis 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) @@ -418,8 +450,7 @@ 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[OF - simple_function_finite simple_function_finite] + unfolding conditional_entropy_eq_ce_with_hypothesis by (simp add: comp_def image_image[symmetric]) finally show ?thesis . qed @@ -433,11 +464,11 @@ unfolding ce_OB_eq_ce_t .. also have "\ = \(t\OB) - \(t\OB | fst)" 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 + by (subst entropy_commute) simp also have "\ \ \(t\OB)" - using conditional_entropy_positive[of "t\OB" fst] by simp + using conditional_entropy_nonneg[of "t\OB" fst] by simp also have "\ \ log b (real (card ((t\OB)`msgs)))" - using entropy_le_card[of "t\OB"] by simp + using entropy_le_card[of "t\OB", OF simple_distributedI[OF simple_function_finite refl]] by simp also have "\ \ log b (real (n + 1)^card observations)" using card_T_bound not_empty by (auto intro!: log_le simp: card_gt_0_iff power_real_of_nat simp del: real_of_nat_power)