# HG changeset patch # User hoelzl # Date 1349868650 -7200 # Node ID 9a2a53be24a2d4bcdcf4aa34f9d142b8ead7a4ee # Parent b1493798d2532cb0e40881c8e2939e759b885b3c# Parent ace9b5a83e603e1d564b8f769301ce8ceb6746ee merged diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -33,15 +33,22 @@ {a \ b | a b. a \ sets A \ b \ sets B} (\X. \\<^isup>+x. (\\<^isup>+y. indicator X (x,y) \B) \A)" +lemma pair_measure_closed: "{a \ b | a b. a \ sets A \ b \ sets B} \ Pow (space A \ space B)" + using space_closed[of A] space_closed[of B] by auto + lemma space_pair_measure: "space (A \\<^isub>M B) = space A \ space B" - unfolding pair_measure_def using space_closed[of A] space_closed[of B] - by (intro space_measure_of) auto + unfolding pair_measure_def using pair_measure_closed[of A B] + by (rule space_measure_of) 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 + unfolding pair_measure_def using pair_measure_closed[of A B] + by (rule sets_measure_of) + +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 pair_measureI[intro, simp]: "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^isub>M B)" @@ -54,21 +61,30 @@ unfolding pair_measure_def using 1 2 by (intro measurable_measure_of) (auto dest: sets_into_space) +lemma measurable_Pair: + assumes f: "f \ measurable M M1" and g: "g \ measurable M M2" + shows "(\x. (f x, g x)) \ measurable M (M1 \\<^isub>M M2)" +proof (rule measurable_pair_measureI) + show "(\x. (f x, g x)) \ space M \ space M1 \ space M2" + using f g by (auto simp: measurable_def) + fix A B assume *: "A \ sets M1" "B \ sets M2" + have "(\x. (f x, g x)) -` (A \ B) \ space M = (f -` A \ space M) \ (g -` B \ space M)" + by auto + also have "\ \ sets M" + by (rule Int) (auto intro!: measurable_sets * f g) + finally show "(\x. (f x, g x)) -` (A \ B) \ space M \ sets M" . +qed + +lemma measurable_pair: + assumes "(fst \ f) \ measurable M M1" "(snd \ f) \ measurable M M2" + shows "f \ measurable M (M1 \\<^isub>M M2)" + using measurable_Pair[OF assms] by simp + 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) + by (auto simp: fst_vimage_eq_Times space_pair_measure sets_into_space times_Int_times measurable_def) 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) + by (auto simp: snd_vimage_eq_Times space_pair_measure sets_into_space times_Int_times measurable_def) 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) @@ -84,55 +100,29 @@ 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 auto + using measurable_pair[of f M M1 M2] 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 measurable_split_conv: + "(\(x, y). f x y) \ measurable A B \ (\x. f (fst x) (snd x)) \ measurable A B" + by (intro arg_cong2[where f="op \"]) auto 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) + by (auto intro!: measurable_Pair simp: measurable_split_conv) 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 + using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def) 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 measurable_ident[intro, simp]: "(\x. x) \ measurable M M" + unfolding measurable_def by auto + 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) + by (auto intro!: measurable_Pair) lemma sets_Pair1: assumes A: "A \ sets (M1 \\<^isub>M M2)" shows "Pair x -` A \ sets M2" proof - @@ -144,8 +134,7 @@ 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) + by (auto intro!: measurable_Pair) lemma sets_Pair2: assumes A: "A \ sets (M1 \\<^isub>M M2)" shows "(\x. (x, y)) -` A \ sets M1" proof - @@ -172,46 +161,115 @@ unfolding Int_stable_def by safe (auto simp add: times_Int_times) -lemma finite_measure_cut_measurable: - assumes "sigma_finite_measure M1" "finite_measure M2" - assumes "Q \ sets (M1 \\<^isub>M M2)" - shows "(\x. emeasure M2 (Pair x -` Q)) \ borel_measurable M1" +lemma (in finite_measure) finite_measure_cut_measurable: + assumes "Q \ sets (N \\<^isub>M M)" + shows "(\x. emeasure M (Pair x -` Q)) \ borel_measurable N" (is "?s Q \ _") + using Int_stable_pair_measure_generator pair_measure_closed assms + unfolding sets_pair_measure +proof (induct rule: sigma_sets_induct_disjoint) + case (compl A) + with sets_into_space have "\x. emeasure M (Pair x -` ((space N \ space M) - A)) = + (if x \ space N then emeasure M (space M) - ?s A x else 0)" + unfolding sets_pair_measure[symmetric] + by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) + with compl top show ?case + by (auto intro!: measurable_If simp: space_pair_measure) +next + case (union F) + moreover then have "\x. emeasure M (\i. Pair x -` F i) = (\i. ?s (F i) x)" + unfolding sets_pair_measure[symmetric] + by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1) + ultimately show ?case + by (auto simp: vimage_UN) +qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If) + +lemma (in sigma_finite_measure) measurable_emeasure_Pair: + assumes Q: "Q \ sets (N \\<^isub>M M)" shows "(\x. emeasure M (Pair x -` Q)) \ borel_measurable N" (is "?s Q \ _") proof - - interpret M1: sigma_finite_measure M1 by fact - interpret M2: finite_measure M2 by fact - 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 - proof (intro dynkin_systemI) - fix A assume "A \ ?D" then show "A \ ?\" - using sets_into_space[of A "M1 \\<^isub>M M2"] by simp - next - from top show "?\ \ ?D" - by (auto simp add: if_distrib intro!: measurable_If) - next - 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 \ ?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) + from sigma_finite_disjoint guess F . note F = this + then have F_sets: "\i. F i \ sets M" by auto + let ?C = "\x i. F i \ Pair x -` Q" + { fix i + have [simp]: "space N \ F i \ space N \ space M = space N \ F i" + using F sets_into_space by auto + let ?R = "density M (indicator (F i))" + have "finite_measure ?R" + using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) + then have "(\x. emeasure ?R (Pair x -` (space N \ space ?R \ Q))) \ borel_measurable N" + by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q) + moreover have "\x. emeasure ?R (Pair x -` (space N \ space ?R \ Q)) + = emeasure M (F i \ Pair x -` (space N \ space ?R \ Q))" + using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) + moreover have "\x. F i \ Pair x -` (space N \ space ?R \ Q) = ?C x i" + using sets_into_space[OF Q] by (auto simp: space_pair_measure) + ultimately have "(\x. emeasure M (?C x i)) \ borel_measurable N" + by simp } + moreover + { fix x + have "(\i. emeasure M (?C x i)) = emeasure M (\i. ?C x i)" + proof (intro suminf_emeasure) + show "range (?C x) \ sets M" + using F `Q \ sets (N \\<^isub>M M)` 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[OF `Q \ sets (N \\<^isub>M M)`] + by (auto simp: space_pair_measure) + finally have "emeasure M (Pair x -` Q) = (\i. emeasure M (?C x i))" + by simp } + ultimately show ?thesis using `Q \ sets (N \\<^isub>M M)` F_sets + by auto +qed + +lemma (in sigma_finite_measure) emeasure_pair_measure: + assumes "X \ sets (N \\<^isub>M M)" + shows "emeasure (N \\<^isub>M M) X = (\\<^isup>+ x. \\<^isup>+ y. indicator X (x, y) \M \N)" (is "_ = ?\ X") +proof (rule emeasure_measure_of[OF pair_measure_def]) + show "positive (sets (N \\<^isub>M M)) ?\" + 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 (N \\<^isub>M M)) ?\" + proof (rule countably_additiveI) + fix F :: "nat \ ('b \ 'a) set" assume F: "range F \ sets (N \\<^isub>M M)" "disjoint_family F" + from F have *: "\i. F i \ sets (N \\<^isub>M M)" "(\i. F i) \ sets (N \\<^isub>M M)" by auto + moreover from F have "\i. (\x. emeasure M (Pair x -` F i)) \ borel_measurable N" + by (intro measurable_emeasure_Pair) 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 M" + 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 - 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 + show "{a \ b |a b. a \ sets N \ b \ sets M} \ Pow (space N \ space M)" + using space_closed[of N] space_closed[of M] by auto +qed fact + +lemma (in sigma_finite_measure) emeasure_pair_measure_alt: + assumes X: "X \ sets (N \\<^isub>M M)" + shows "emeasure (N \\<^isub>M M) X = (\\<^isup>+x. emeasure M (Pair x -` X) \N)" +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 sigma_finite_measure) emeasure_pair_measure_Times: + assumes A: "A \ sets N" and B: "B \ sets M" + shows "emeasure (N \\<^isub>M M) (A \ B) = emeasure N A * emeasure M B" +proof - + have "emeasure (N \\<^isub>M M) (A \ B) = (\\<^isup>+x. emeasure M B * indicator A x \N)" + using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt) + also have "\ = emeasure M B * emeasure N A" + using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator) + finally show ?thesis + by (simp add: ac_simps) qed subsection {* Binary products of $\sigma$-finite emeasure spaces *} @@ -219,114 +277,21 @@ locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 for M1 :: "'a measure" and M2 :: "'b measure" -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 - - 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 - have [simp]: "space M1 \ F i \ space M1 \ space M2 = space M1 \ F i" - 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 ?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. 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. 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 (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[OF `Q \ sets (M1 \\<^isub>M M2)`] - by (auto simp: space_pair_measure) - finally have "emeasure M2 (Pair x -` Q) = (\i. emeasure M2 (?C x i))" - by simp } - ultimately show ?thesis using `Q \ sets (M1 \\<^isub>M M2)` F_sets - by auto -qed + "Q \ sets (M1 \\<^isub>M M2) \ (\x. emeasure M2 (Pair x -` Q)) \ borel_measurable M1" + using M2.measurable_emeasure_Pair . 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 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] + note M1.measurable_emeasure_Pair[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) 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) 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 "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 pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: 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 \ @@ -366,7 +331,7 @@ qed qed -sublocale pair_sigma_finite \ sigma_finite_measure "M1 \\<^isub>M M2" +sublocale pair_sigma_finite \ P: 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) \ \)" @@ -396,7 +361,6 @@ 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 let ?E = "{a \ b |a b. a \ sets M1 \ b \ sets M2}" show ?thesis @@ -408,7 +372,7 @@ 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) \ \" + show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" using F by (auto simp: space_pair_measure) next fix X assume "X \ ?E" @@ -416,7 +380,7 @@ 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 + by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr measurable_pair_swap' ac_simps) qed qed @@ -426,138 +390,14 @@ shows "emeasure (M1 \\<^isub>M M2) A = (\\<^isup>+y. emeasure M1 ((\x. (x, y)) -` A) \M2)" (is "_ = ?\ A") proof - - 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]) + M1.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 (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 (M1 \\<^isub>M M2) f" -proof - - 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 (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 (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 (rule sets_Pair1[OF measurable_sets]) auto - ultimately have "simple_function M2 (\ y. f (x, y))" - apply (rule_tac simple_function_cong[THEN iffD2, OF _]) - apply (rule simple_function_indicator_representation[OF f(1)]) - 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 (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 (M1 \\<^isub>M M2))" by (rule simple_functionD) - next - 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 (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) * 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 (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 (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] - emeasure_pair_measure_alt[symmetric]) -qed - -lemma (in pair_sigma_finite) positive_integral_fst_measurable: - 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 (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 (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_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 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) \(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 (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 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!: 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) 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 - 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) \(M2 \\<^isub>M M1))" - by simp - 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 (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 (M1 \\<^isub>M M2). Q x" shows "AE x in M1. (AE y in M2. Q (x, y))" @@ -568,7 +408,7 @@ 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) + by (auto simp: M2.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" @@ -593,7 +433,7 @@ 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) + by (simp add: M2.emeasure_pair_measure) also have "\ = (\\<^isup>+ x. \\<^isup>+ y. 0 \M2 \M1)" using ae apply (safe intro!: positive_integral_cong_AE) @@ -609,24 +449,6 @@ (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)" @@ -648,6 +470,82 @@ done qed +section "Fubinis theorem" + +lemma measurable_compose_Pair1: + "x \ space M1 \ g \ measurable (M1 \\<^isub>M M2) L \ (\y. g (x, y)) \ measurable M2 L" + by (rule measurable_compose[OF measurable_Pair]) auto + +lemma (in pair_sigma_finite) borel_measurable_positive_integral_fst: + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" "\x. 0 \ f x" + shows "(\x. \\<^isup>+ y. f (x, y) \M2) \ borel_measurable M1" +using f proof induct + case (cong u v) + then have "\w x. w \ space M1 \ x \ space M2 \ u (w, x) = v (w, x)" + by (auto simp: space_pair_measure) + show ?case + apply (subst measurable_cong) + apply (rule positive_integral_cong) + apply fact+ + done +next + case (set Q) + have [simp]: "\x y. indicator Q (x, y) = indicator (Pair x -` Q) y" + by (auto simp: indicator_def) + have "\x. x \ space M1 \ emeasure M2 (Pair x -` Q) = \\<^isup>+ y. indicator Q (x, y) \M2" + by (simp add: sets_Pair1[OF set]) + from this M2.measurable_emeasure_Pair[OF set] show ?case + by (rule measurable_cong[THEN iffD1]) +qed (simp_all add: positive_integral_add positive_integral_cmult measurable_compose_Pair1 + positive_integral_monotone_convergence_SUP incseq_def le_fun_def + cong: measurable_cong) + +lemma (in pair_sigma_finite) positive_integral_fst: + assumes f: "f \ borel_measurable (M1 \\<^isub>M M2)" "\x. 0 \ f x" + shows "(\\<^isup>+ x. \\<^isup>+ y. f (x, y) \M2 \M1) = integral\<^isup>P (M1 \\<^isub>M M2) f" (is "?I f = _") +using f proof induct + case (cong u v) + moreover then have "?I u = ?I v" + by (intro positive_integral_cong) (auto simp: space_pair_measure) + ultimately show ?case + by (simp cong: positive_integral_cong) +qed (simp_all add: M2.emeasure_pair_measure positive_integral_cmult positive_integral_add + positive_integral_monotone_convergence_SUP + measurable_compose_Pair1 positive_integral_positive + borel_measurable_positive_integral_fst positive_integral_mono incseq_def le_fun_def + cong: positive_integral_cong) + +lemma (in pair_sigma_finite) positive_integral_fst_measurable: + 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 (M1 \\<^isub>M M2) f" + using f + borel_measurable_positive_integral_fst[of "\x. max 0 (f x)"] + positive_integral_fst[of "\x. max 0 (f x)"] + unfolding positive_integral_max_0 by auto + +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 + 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) \(M2 \\<^isub>M M1))" + by simp + 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 (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) integrable_product_swap: assumes "integrable (M1 \\<^isub>M M2) f" shows "integrable (M2 \\<^isub>M M1) (\(x,y). f (y,x))" @@ -817,7 +715,7 @@ 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 B.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) @@ -861,7 +759,7 @@ by simp show "emeasure ?L A = emeasure ?R A" - apply (subst L.emeasure_pair_measure[OF A]) + apply (subst D2.emeasure_pair_measure[OF A]) apply (subst emeasure_density) using f_fst g_snd apply (simp add: split_beta') using A apply simp @@ -924,7 +822,7 @@ 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) + by (rule T.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)" @@ -933,7 +831,7 @@ 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]) + using fg by (intro N.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 diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Wed Oct 10 13:30:50 2012 +0200 @@ -632,7 +632,7 @@ "(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: +lemma borel_measurable_euclidean_component': "(\x::'a::euclidean_space. x $$ i) \ borel_measurable borel" proof (rule borel_measurableI) fix S::"real set" assume "open S" @@ -641,13 +641,18 @@ by (auto intro: borel_open) qed +lemma borel_measurable_euclidean_component: + fixes f :: "'a \ 'b::euclidean_space" + assumes f: "f \ borel_measurable M" + shows "(\x. f x $$ i) \ borel_measurable M" + using measurable_comp[OF f borel_measurable_euclidean_component'] by (simp add: comp_def) + 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 fix i assume "f \ borel_measurable M" then show "(\x. f x $$ i) \ borel_measurable M" - using measurable_comp[of f _ _ "\x. x $$ i", unfolded comp_def] by (auto intro: borel_measurable_euclidean_component) next assume f: "\ix. f x $$ i) \ borel_measurable M" @@ -657,6 +662,144 @@ subsection "Borel measurable operators" +lemma borel_measurable_continuous_on1: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes "continuous_on UNIV f" + shows "f \ borel_measurable borel" + 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::topological_space" + assumes f: "continuous_on UNIV f" and g: "g \ borel_measurable M" + shows "(\x. f (g x)) \ borel_measurable M" + using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def) + +lemma borel_measurable_continuous_on_open': + 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_measurableI) + fix S :: "'b set" assume "open S" + then have "open {x\A. f x \ S}" + by (intro continuous_open_preimage[OF cont]) auto + then have *: "{x\A. f x \ S} \ sets borel" by auto + have "?f -` S \ space borel = + {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!: Un) + finally show "?f -` S \ space borel \ sets borel" . +qed + +lemma borel_measurable_continuous_on_open: + fixes f :: "'a::topological_space \ 'b::t1_space" + assumes cont: "continuous_on A f" "open A" + assumes g: "g \ borel_measurable M" + shows "(\x. if g x \ A then f (g x) else c) \ borel_measurable M" + using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c] + by (simp add: comp_def) + +lemma borel_measurable_uminus[simp, intro]: + fixes g :: "'a \ real" + assumes g: "g \ borel_measurable M" + shows "(\x. - g x) \ borel_measurable M" + by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id) + +lemma euclidean_component_prod: + fixes x :: "'a :: euclidean_space \ 'b :: euclidean_space" + shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))" + unfolding euclidean_component_def basis_prod_def inner_prod_def by auto + +lemma borel_measurable_Pair[simp, intro]: + fixes f :: "'a \ 'b::ordered_euclidean_space" and g :: "'a \ 'c::ordered_euclidean_space" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows "(\x. (f x, g x)) \ borel_measurable M" +proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI) + fix i and a :: real assume i: "i < DIM('b \ 'c)" + have [simp]: "\P A B C. {w. (P \ A w \ B w) \ (\ P \ A w \ C w)} = + {w. A w \ (P \ B w) \ (\ P \ C w)}" by auto + from i f g show "{w \ space M. (f w, g w) $$ i \ a} \ sets M" + by (auto simp: euclidean_component_prod intro!: sets_Collect borel_measurable_euclidean_component) +qed + +lemma continuous_on_fst: "continuous_on UNIV fst" +proof - + have [simp]: "range fst = UNIV" by (auto simp: image_iff) + show ?thesis + using closed_vimage_fst + by (auto simp: continuous_on_closed closed_closedin vimage_def) +qed + +lemma continuous_on_snd: "continuous_on UNIV snd" +proof - + have [simp]: "range snd = UNIV" by (auto simp: image_iff) + show ?thesis + using closed_vimage_snd + by (auto simp: continuous_on_closed closed_closedin vimage_def) +qed + +lemma borel_measurable_continuous_Pair: + fixes f :: "'a \ 'b::ordered_euclidean_space" and g :: "'a \ 'c::ordered_euclidean_space" + assumes [simp]: "f \ borel_measurable M" + assumes [simp]: "g \ borel_measurable M" + assumes H: "continuous_on UNIV (\x. H (fst x) (snd x))" + shows "(\x. H (f x) (g x)) \ borel_measurable M" +proof - + have eq: "(\x. H (f x) (g x)) = (\x. (\x. H (fst x) (snd x)) (f x, g x))" by auto + show ?thesis + unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto +qed + +lemma borel_measurable_add[simp, intro]: + fixes f g :: "'a \ 'c::ordered_euclidean_space" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows "(\x. f x + g x) \ borel_measurable M" + using f g + by (rule borel_measurable_continuous_Pair) + (auto intro: continuous_on_fst continuous_on_snd continuous_on_add) + +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" +proof cases + assume "finite S" + thus ?thesis using assms by induct auto +qed simp + +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 borel_measurable_times[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" + using f g + by (rule borel_measurable_continuous_Pair) + (auto intro: continuous_on_fst continuous_on_snd continuous_on_mult) + +lemma continuous_on_dist: + fixes f :: "'a :: t2_space \ 'b :: metric_space" + shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. dist (f x) (g x))" + unfolding continuous_on_eq_continuous_within by (auto simp: continuous_dist) + +lemma borel_measurable_dist[simp, intro]: + fixes g f :: "'a \ 'b::ordered_euclidean_space" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows "(\x. dist (f x) (g x)) \ borel_measurable M" + using f g + by (rule borel_measurable_continuous_Pair) + (intro continuous_on_dist continuous_on_fst continuous_on_snd) + lemma affine_borel_measurable_vector: fixes f :: "'a \ 'x::real_normed_vector" assumes "f \ borel_measurable M" @@ -683,116 +826,6 @@ shows "(\x. a + (g x) * b) \ borel_measurable M" using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute) -lemma borel_measurable_add[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" -proof - - have 1: "\a. {w\space M. a \ f w + g w} = {w \ space M. a + g w * -1 \ f w}" - by auto - have "\a. (\w. a + (g w) * -1) \ borel_measurable M" - by (rule affine_borel_measurable [OF g]) - then have "\a. {w \ space M. (\w. a + (g w) * -1)(w) \ f w} \ sets M" using f - by auto - then have "\a. {w \ space M. a \ f w + g w} \ sets M" - by (simp add: 1) - then show ?thesis - by (simp add: borel_measurable_iff_ge) -qed - -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" -proof cases - assume "finite S" - thus ?thesis using assms by induct auto -qed simp - -lemma borel_measurable_square: - fixes f :: "'a \ real" - assumes f: "f \ borel_measurable M" - shows "(\x. (f x)^2) \ borel_measurable M" -proof - - { - fix a - have "{w \ space M. (f w)\ \ a} \ sets M" - proof (cases rule: linorder_cases [of a 0]) - case less - hence "{w \ space M. (f w)\ \ a} = {}" - by auto (metis less order_le_less_trans power2_less_0) - also have "... \ sets M" - by (rule empty_sets) - finally show ?thesis . - next - case equal - hence "{w \ space M. (f w)\ \ a} = - {w \ space M. f w \ 0} \ {w \ space M. 0 \ f w}" - by auto - also have "... \ sets M" - apply (insert f) - apply (rule Int) - apply (simp add: borel_measurable_iff_le) - apply (simp add: borel_measurable_iff_ge) - done - finally show ?thesis . - next - case greater - have "\x. (f x ^ 2 \ sqrt a ^ 2) = (- sqrt a \ f x & f x \ sqrt a)" - by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs - real_sqrt_le_iff real_sqrt_power) - hence "{w \ space M. (f w)\ \ a} = - {w \ space M. -(sqrt a) \ f w} \ {w \ space M. f w \ sqrt a}" - using greater by auto - also have "... \ sets M" - apply (insert f) - apply (rule Int) - apply (simp add: borel_measurable_iff_ge) - apply (simp add: borel_measurable_iff_le) - done - finally show ?thesis . - qed - } - thus ?thesis by (auto simp add: borel_measurable_iff_le) -qed - -lemma times_eq_sum_squares: - fixes x::real - shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4" -by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric]) - -lemma borel_measurable_uminus[simp, intro]: - fixes g :: "'a \ real" - assumes g: "g \ borel_measurable M" - shows "(\x. - g x) \ borel_measurable M" -proof - - have "(\x. - g x) = (\x. 0 + (g x) * -1)" - by simp - also have "... \ borel_measurable M" - by (fast intro: affine_borel_measurable g) - finally show ?thesis . -qed - -lemma borel_measurable_times[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" -proof - - have 1: "(\x. 0 + (f x + g x)\ * inverse 4) \ borel_measurable M" - using assms by (fast intro: affine_borel_measurable borel_measurable_square) - have "(\x. -((f x + -g x) ^ 2 * inverse 4)) = - (\x. 0 + ((f x + -g x) ^ 2 * inverse -4))" - by (simp add: minus_divide_right) - also have "... \ borel_measurable M" - using f g by (fast intro: affine_borel_measurable borel_measurable_square f g) - finally have 2: "(\x. -((f x + -g x) ^ 2 * inverse 4)) \ borel_measurable M" . - show ?thesis - apply (simp add: times_eq_sum_squares diff_minus) - using 1 2 by simp -qed - lemma borel_measurable_setprod[simp, intro]: fixes f :: "'c \ 'a \ real" assumes "\i. i \ S \ f i \ borel_measurable M" @@ -802,26 +835,17 @@ thus ?thesis using assms by induct auto qed simp -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 borel_measurable_inverse[simp, intro]: fixes f :: "'a \ real" - assumes "f \ borel_measurable M" + assumes f: "f \ borel_measurable M" shows "(\x. inverse (f x)) \ borel_measurable M" - unfolding borel_measurable_iff_ge unfolding inverse_eq_divide -proof safe - fix a :: real - have *: "{w \ space M. a \ 1 / f w} = - ({w \ space M. 0 < f w} \ {w \ space M. a * f w \ 1}) \ - ({w \ space M. f w < 0} \ {w \ space M. 1 \ a * f w}) \ - ({w \ space M. f w = 0} \ {w \ space M. a \ 0})" by (auto simp: le_divide_eq) - show "{w \ space M. a \ 1 / f w} \ sets M" using assms unfolding * - by (auto intro!: Int Un) +proof - + have *: "\x::real. inverse x = (if x \ UNIV - {0} then inverse x else 0)" by auto + show ?thesis + apply (subst *) + apply (rule borel_measurable_continuous_on_open) + apply (auto intro!: f continuous_on_inverse continuous_on_id) + done qed lemma borel_measurable_divide[simp, intro]: @@ -837,30 +861,14 @@ assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. max (g x) (f x)) \ borel_measurable M" - unfolding borel_measurable_iff_le -proof safe - fix a - have "{x \ space M. max (g x) (f x) \ a} = - {x \ space M. g x \ a} \ {x \ space M. f x \ a}" by auto - thus "{x \ space M. max (g x) (f x) \ a} \ sets M" - using assms unfolding borel_measurable_iff_le - by (auto intro!: Int) -qed + unfolding max_def by (auto intro!: assms measurable_If) lemma borel_measurable_min[intro, simp]: fixes f g :: "'a \ real" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. min (g x) (f x)) \ borel_measurable M" - unfolding borel_measurable_iff_ge -proof safe - fix a - have "{x \ space M. a \ min (g x) (f x)} = - {x \ space M. a \ g x} \ {x \ space M. a \ f x}" by auto - thus "{x \ space M. a \ min (g x) (f x)} \ sets M" - using assms unfolding borel_measurable_iff_ge - by (auto intro!: Int) -qed + unfolding min_def by (auto intro!: assms measurable_If) lemma borel_measurable_abs[simp, intro]: assumes "f \ borel_measurable M" @@ -872,76 +880,50 @@ lemma borel_measurable_nth[simp, intro]: "(\x::real^'n. x $ i) \ borel_measurable borel" - using borel_measurable_euclidean_component + using borel_measurable_euclidean_component' unfolding nth_conv_component by auto -lemma borel_measurable_continuous_on1: - fixes f :: "'a::topological_space \ 'b::t1_space" - assumes "continuous_on UNIV f" - shows "f \ borel_measurable borel" - 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_measurableI) - fix S :: "'b set" assume "open S" - then have "open {x\A. f x \ S}" - by (intro continuous_open_preimage[OF cont]) auto - then have *: "{x\A. f x \ S} \ sets borel" by auto - have "?f -` S \ space borel = - {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!: Un) - finally show "?f -` S \ space borel \ sets borel" . -qed - 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" - shows "q \ X \ borel_measurable M" + shows "(\x. q (X x)) \ borel_measurable M" proof - - have "(\x. if x \ {a <..< b} then q x else 0) \ borel_measurable borel" - proof (rule borel_measurable_continuous_on) + have "(\x. if X x \ {a <..< b} then q (X x) else 0) \ borel_measurable M" (is "?qX") + proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)]) show "open {a<..x. if x \ {a <..< b} then q x else 0) \ X \ borel_measurable M" (is ?qX) - using X by (intro measurable_comp) auto - moreover have "?qX \ q \ X \ borel_measurable M" + moreover have "?qX \ (\x. q (X x)) \ borel_measurable M" using X by (intro measurable_cong) auto ultimately show ?thesis by simp qed -lemma borel_measurable_borel_log: assumes "1 < b" shows "log b \ borel_measurable borel" +lemma borel_measurable_ln[simp,intro]: + assumes f: "f \ borel_measurable M" + shows "(\x. ln (f x)) \ borel_measurable M" proof - { fix x :: real assume x: "x \ 0" { fix x::real assume "x \ 0" then have "\u. exp u = x \ False" by auto } - from this[of x] x this[of 0] have "log b 0 = log b x" - by (auto simp: ln_def log_def) } - note log_imp = this - have "(\x. if x \ {0<..} then log b x else log b 0) \ borel_measurable borel" - proof (rule borel_measurable_continuous_on) - show "continuous_on {0<..} (log b)" - by (auto intro!: continuous_at_imp_continuous_on DERIV_log DERIV_isCont + from this[of x] x this[of 0] have "ln 0 = ln x" + by (auto simp: ln_def) } + note ln_imp = this + have "(\x. if f x \ {0<..} then ln (f x) else ln 0) \ borel_measurable M" + proof (rule borel_measurable_continuous_on_open[OF _ _ f]) + show "continuous_on {0<..} ln" + by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont simp: continuous_isCont[symmetric]) show "open ({0<..}::real set)" by auto qed - also have "(\x. if x \ {0<..} then log b x else log b 0) = log b" - by (simp add: fun_eq_iff not_less log_imp) + also have "(\x. if x \ {0<..} then ln x else ln 0) = ln" + by (simp add: fun_eq_iff not_less ln_imp) finally show ?thesis . qed 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`]] - by (simp add: comp_def) + "f \ borel_measurable M \ (\x. log b (f x)) \ borel_measurable M" + unfolding log_def by auto lemma borel_measurable_real_floor: "(\x::real. real \x\) \ borel_measurable borel" @@ -967,45 +949,91 @@ subsection "Borel space on the extended reals" -lemma borel_measurable_ereal_borel: - "ereal \ borel_measurable borel" -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 - 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 . + using continuous_on_ereal f by (rule borel_measurable_continuous_on) -lemma borel_measurable_real_of_ereal_borel: - "(real :: ereal \ real) \ borel_measurable borel" -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 - show "(real -` B \ space borel :: ereal set) \ sets borel" - proof cases - assume "0 \ B" - then have *: "real -` B = real -` (B - {0}) \ {-\, \, 0::ereal}" - by (auto simp add: real_of_ereal_eq_0) - then show "(real -` B :: ereal set) \ space borel \ sets borel" - using open_real by auto - next - assume "0 \ B" - then have *: "(real -` B :: ereal set) = real -` (B - {0})" - by (auto simp add: real_of_ereal_eq_0) - then show "(real -` B :: ereal set) \ space borel \ sets borel" - using open_real by auto - qed +lemma borel_measurable_real_of_ereal[simp, intro]: + fixes f :: "'a \ ereal" + assumes f: "f \ borel_measurable M" + shows "(\x. real (f x)) \ borel_measurable M" +proof - + have "(\x. if f x \ UNIV - { \, - \ } then real (f x) else 0) \ borel_measurable M" + using continuous_on_real + by (rule borel_measurable_continuous_on_open[OF _ _ f]) auto + also have "(\x. if f x \ UNIV - { \, - \ } then real (f x) else 0) = (\x. real (f x))" + by auto + finally show ?thesis . +qed + +lemma borel_measurable_ereal_cases: + fixes f :: "'a \ ereal" + assumes f: "f \ borel_measurable M" + assumes H: "(\x. H (ereal (real (f x)))) \ borel_measurable M" + shows "(\x. H (f x)) \ borel_measurable M" +proof - + let ?F = "\x. if x \ f -` {\} then H \ else if x \ f -` {-\} then H (-\) else H (ereal (real (f x)))" + { fix x have "H (f x) = ?F x" by (cases "f x") auto } + moreover + have "?F \ borel_measurable M" + by (intro measurable_If_set f measurable_sets[OF f] H) auto + ultimately + show ?thesis by simp qed -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 + fixes f :: "'a \ ereal" assumes f[simp]: "f \ borel_measurable M" + shows borel_measurable_ereal_abs[intro, simp]: "(\x. \f x\) \ borel_measurable M" + and borel_measurable_ereal_inverse[simp, intro]: "(\x. inverse (f x) :: ereal) \ borel_measurable M" + and borel_measurable_uminus_ereal[intro]: "(\x. - f x :: ereal) \ borel_measurable M" + by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) + +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 sets_Collect_If_set: + assumes "A \ space M \ sets M" "{x \ space M. P x} \ sets M" "{x \ space M. Q x} \ sets M" + shows "{x \ space M. if x \ A then P x else Q x} \ sets M" +proof - + have *: "{x \ space M. if x \ A then P x else Q x} = + {x \ space M. if x \ A \ space M then P x else Q x}" by auto + show ?thesis unfolding * unfolding if_bool_eq_conj using assms + by (auto intro!: sets_Collect simp: Int_def conj_commute) +qed + +lemma set_Collect_ereal2: + fixes f g :: "'a \ ereal" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + assumes H: "{x \ space M. H (ereal (real (f x))) (ereal (real (g x)))} \ sets M" + "{x \ space M. H (-\) (ereal (real (g x)))} \ sets M" + "{x \ space M. H (\) (ereal (real (g x)))} \ sets M" + "{x \ space M. H (ereal (real (f x))) (-\)} \ sets M" + "{x \ space M. H (ereal (real (f x))) (\)} \ sets M" + shows "{x \ space M. H (f x) (g x)} \ sets M" +proof - + let ?G = "\y x. if x \ g -` {\} then H y \ else if x \ g -` {-\} then H y (-\) else H y (ereal (real (g x)))" + let ?F = "\x. if x \ f -` {\} then ?G \ x else if x \ f -` {-\} then ?G (-\) x else ?G (ereal (real (f x))) x" + { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } + moreover + have "{x \ space M. ?F x} \ sets M" + by (intro sets_Collect H measurable_sets[OF f] measurable_sets[OF g] sets_Collect_If_set) auto + ultimately + show ?thesis by simp +qed + +lemma + fixes f g :: "'a \ ereal" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows borel_measurable_ereal_le[intro,simp]: "{x \ space M. f x \ g x} \ sets M" + and borel_measurable_ereal_less[intro,simp]: "{x \ space M. f x < g x} \ sets M" + and borel_measurable_ereal_eq[intro,simp]: "{w \ space M. f w = g w} \ sets M" + and borel_measurable_ereal_neq[intro,simp]: "{w \ space M. f w \ g w} \ sets M" + using f g by (auto simp: f g set_Collect_ereal2[OF f g] intro!: sets_Collect_neg) lemma borel_measurable_ereal_iff: shows "(\x. ereal (f x)) \ borel_measurable M \ f \ borel_measurable M" @@ -1029,52 +1057,6 @@ finally show "f \ borel_measurable M" . qed (auto intro: measurable_sets borel_measurable_real_of_ereal) -lemma less_eq_ge_measurable: - fixes f :: "'a \ 'c::linorder" - shows "f -` {a <..} \ space M \ sets M \ f -` {..a} \ space M \ sets M" -proof - assume "f -` {a <..} \ space M \ sets M" - moreover have "f -` {..a} \ space M = space M - f -` {a <..} \ space M" by auto - ultimately show "f -` {..a} \ space M \ sets M" by auto -next - assume "f -` {..a} \ space M \ sets M" - moreover have "f -` {a <..} \ space M = space M - f -` {..a} \ space M" by auto - ultimately show "f -` {a <..} \ space M \ sets M" by auto -qed - -lemma greater_eq_le_measurable: - fixes f :: "'a \ 'c::linorder" - shows "f -` {..< a} \ space M \ sets M \ f -` {a ..} \ space M \ sets M" -proof - assume "f -` {a ..} \ space M \ sets M" - moreover have "f -` {..< a} \ space M = space M - f -` {a ..} \ space M" by auto - ultimately show "f -` {..< a} \ space M \ sets M" by auto -next - assume "f -` {..< a} \ space M \ sets M" - moreover have "f -` {a ..} \ space M = space M - f -` {..< a} \ space M" by auto - ultimately show "f -` {a ..} \ space M \ sets M" by auto -qed - -lemma borel_measurable_uminus_borel_ereal: - "(uminus :: ereal \ ereal) \ borel_measurable borel" -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 - -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 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 borel_measurable_eq_atMost_ereal: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" @@ -1118,94 +1100,88 @@ then show "f \ borel_measurable M" by simp qed (simp add: measurable_sets) +lemma greater_eq_le_measurable: + fixes f :: "'a \ 'c::linorder" + shows "f -` {..< a} \ space M \ sets M \ f -` {a ..} \ space M \ sets M" +proof + assume "f -` {a ..} \ space M \ sets M" + moreover have "f -` {..< a} \ space M = space M - f -` {a ..} \ space M" by auto + ultimately show "f -` {..< a} \ space M \ sets M" by auto +next + assume "f -` {..< a} \ space M \ sets M" + moreover have "f -` {a ..} \ space M = space M - f -` {..< a} \ space M" by auto + ultimately show "f -` {a ..} \ space M \ sets M" by auto +qed + 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 less_eq_ge_measurable: + fixes f :: "'a \ 'c::linorder" + shows "f -` {a <..} \ space M \ sets M \ f -` {..a} \ space M \ sets M" +proof + assume "f -` {a <..} \ space M \ sets M" + moreover have "f -` {..a} \ space M = space M - f -` {a <..} \ space M" by auto + ultimately show "f -` {..a} \ space M \ sets M" by auto +next + assume "f -` {..a} \ space M \ sets M" + moreover have "f -` {a <..} \ space M = space M - f -` {..a} \ space M" by auto + ultimately show "f -` {a <..} \ space M \ sets M" by auto +qed + 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 borel_measurable_ereal_eq_const: - fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" - shows "{x\space M. f x = c} \ sets M" -proof - - have "{x\space M. f x = c} = (f -` {c} \ space M)" by auto - then show ?thesis using assms by (auto intro!: measurable_sets) -qed - -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 - - have "{x\space M. f x \ c} = space M - (f -` {c} \ space M)" by auto - then show ?thesis using assms by (auto intro!: measurable_sets) -qed - -lemma borel_measurable_ereal_le[intro,simp]: - fixes f g :: "'a \ ereal" +lemma borel_measurable_ereal2: + fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" - shows "{x \ space M. f x \ g x} \ sets M" + assumes H: "(\x. H (ereal (real (f x))) (ereal (real (g x)))) \ borel_measurable M" + "(\x. H (-\) (ereal (real (g x)))) \ borel_measurable M" + "(\x. H (\) (ereal (real (g x)))) \ borel_measurable M" + "(\x. H (ereal (real (f x))) (-\)) \ borel_measurable M" + "(\x. H (ereal (real (f x))) (\)) \ borel_measurable M" + shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - - have "{x \ space M. f x \ g x} = - {x \ space M. real (f x) \ real (g x)} - (f -` {\, -\} \ space M \ g -` {\, -\} \ space M) \ - f -` {-\} \ space M \ g -` {\} \ space M" (is "?l = ?r") - proof (intro set_eqI) - fix x show "x \ ?l \ x \ ?r" by (cases rule: ereal2_cases[of "f x" "g x"]) auto - qed - with f g show ?thesis by (auto intro!: Un simp: measurable_sets) + let ?G = "\y x. if x \ g -` {\} then H y \ else if x \ g -` {-\} then H y (-\) else H y (ereal (real (g x)))" + let ?F = "\x. if x \ f -` {\} then ?G \ x else if x \ f -` {-\} then ?G (-\) x else ?G (ereal (real (f x))) x" + { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } + moreover + have "?F \ borel_measurable M" + by (intro measurable_If_set measurable_sets[OF f] measurable_sets[OF g] H) auto + ultimately + show ?thesis by simp qed -lemma borel_measurable_ereal_less[intro,simp]: - fixes f :: "'a \ ereal" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "{x \ space M. f x < g x} \ sets M" -proof - - have "{x \ space M. f x < g x} = space M - {x \ space M. g x \ f x}" by auto - then show ?thesis using g f by auto -qed - -lemma borel_measurable_ereal_eq[intro,simp]: - fixes f :: "'a \ ereal" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "{w \ space M. f w = g w} \ sets M" -proof - - have "{x \ space M. f x = g x} = {x \ space M. g x \ f x} \ {x \ space M. f x \ g x}" by auto - then show ?thesis using g f by auto -qed - -lemma borel_measurable_ereal_neq[intro,simp]: - fixes f :: "'a \ ereal" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "{w \ space M. f w \ g w} \ sets M" -proof - - have "{w \ space M. f w \ g w} = space M - {w \ space M. f w = g w}" by auto - thus ?thesis using f g by auto -qed +lemma + fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" + shows borel_measurable_ereal_eq_const: "{x\space M. f x = c} \ sets M" + and borel_measurable_ereal_neq_const: "{x\space M. f x \ c} \ sets M" + using f by auto 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 borel_measurable_ereal_add[intro, simp]: +lemma fixes f :: "'a \ ereal" - assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "(\x. f x + g x) \ borel_measurable M" -proof - - { fix x assume "x \ space M" then have "f x + g x = - (if f x = \ \ g x = \ then \ - else if f x = -\ \ g x = -\ then -\ - else ereal (real (f x) + real (g x)))" - by (cases rule: ereal2_cases[of "f x" "g x"]) auto } - with assms show ?thesis - by (auto cong: measurable_cong simp: split_sets - intro!: Un measurable_If measurable_sets) -qed + assumes [simp]: "f \ borel_measurable M" "g \ borel_measurable M" + shows borel_measurable_ereal_add[intro, simp]: "(\x. f x + g x) \ borel_measurable M" + and borel_measurable_ereal_times[intro, simp]: "(\x. f x * g x) \ borel_measurable M" + and borel_measurable_ereal_min[simp, intro]: "(\x. min (g x) (f x)) \ borel_measurable M" + and borel_measurable_ereal_max[simp, intro]: "(\x. max (g x) (f x)) \ borel_measurable M" + by (auto simp add: borel_measurable_ereal2 measurable_If min_def max_def) + +lemma + fixes f g :: "'a \ ereal" + assumes "f \ borel_measurable M" + assumes "g \ borel_measurable M" + shows borel_measurable_ereal_diff[simp, intro]: "(\x. f x - g x) \ borel_measurable M" + and borel_measurable_ereal_divide[simp, intro]: "(\x. f x / g x) \ borel_measurable M" + unfolding minus_ereal_def divide_ereal_def using assms by auto lemma borel_measurable_ereal_setsum[simp, intro]: fixes f :: "'c \ 'a \ ereal" @@ -1215,39 +1191,7 @@ assume "finite S" thus ?thesis using assms by induct auto -qed (simp add: borel_measurable_const) - -lemma borel_measurable_ereal_abs[intro, simp]: - fixes f :: "'a \ ereal" assumes "f \ borel_measurable M" - shows "(\x. \f x\) \ borel_measurable M" -proof - - { fix x have "\f x\ = (if 0 \ f x then f x else - f x)" by auto } - then show ?thesis using assms by (auto intro!: measurable_If) -qed - -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 - - { fix f g :: "'a \ ereal" - assume b: "f \ borel_measurable M" "g \ borel_measurable M" - and pos: "\x. 0 \ f x" "\x. 0 \ g x" - { fix x have *: "f x * g x = (if f x = 0 \ g x = 0 then 0 - else if f x = \ \ g x = \ then \ - else ereal (real (f x) * real (g x)))" - apply (cases rule: ereal2_cases[of "f x" "g x"]) - using pos[of x] by auto } - with b have "(\x. f x * g x) \ borel_measurable M" - by (auto cong: measurable_cong simp: split_sets - intro!: Un measurable_If measurable_sets) } - note pos_times = this - have *: "(\x. f x * g x) = - (\x. if 0 \ f x \ 0 \ g x \ f x < 0 \ g x < 0 then \f x\ * \g x\ else - (\f x\ * \g x\))" - by (auto simp: fun_eq_iff) - show ?thesis using assms unfolding * - by (intro measurable_If pos_times borel_measurable_uminus_ereal) - (auto simp: split_sets intro!: Int) -qed +qed simp lemma borel_measurable_ereal_setprod[simp, intro]: fixes f :: "'c \ 'a \ ereal" @@ -1258,20 +1202,6 @@ thus ?thesis using assms by induct auto qed simp -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 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 borel_measurable_SUP[simp, intro]: fixes f :: "'d\countable \ 'a \ ereal" assumes "\i. i \ A \ f i \ borel_measurable M" @@ -1298,38 +1228,24 @@ using assms by auto qed -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 borel_measurable_limsup[simp, intro]: +lemma 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 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 + shows borel_measurable_liminf[simp, intro]: "(\x. liminf (\i. f i x)) \ borel_measurable M" + and borel_measurable_limsup[simp, intro]: "(\x. limsup (\i. f i x)) \ borel_measurable M" + unfolding liminf_SUPR_INFI limsup_INFI_SUPR using assms by auto -lemma borel_measurable_ereal_inverse[simp, intro]: - assumes f: "f \ borel_measurable M" shows "(\x. inverse (f x) :: ereal) \ borel_measurable M" +lemma borel_measurable_ereal_LIMSEQ: + fixes u :: "nat \ 'a \ ereal" + assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" + and u: "\i. u i \ borel_measurable M" + shows "u' \ 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) + have "\x. x \ space M \ u' x = liminf (\n. u n x)" + using u' by (simp add: lim_imp_Liminf[symmetric]) + then show ?thesis by (simp add: u cong: measurable_cong) 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" @@ -1354,4 +1270,38 @@ ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) qed -end +lemma sets_Collect_Cauchy: + fixes f :: "nat \ 'a => real" + assumes f: "\i. f i \ borel_measurable M" + shows "{x\space M. Cauchy (\i. f i x)} \ sets M" + unfolding Cauchy_iff2 using f by (auto intro!: sets_Collect) + +lemma borel_measurable_lim: + fixes f :: "nat \ 'a \ real" + assumes f: "\i. f i \ borel_measurable M" + shows "(\x. lim (\i. f i x)) \ borel_measurable M" +proof - + have *: "\x. lim (\i. f i x) = + (if Cauchy (\i. f i x) then lim (\i. if Cauchy (\i. f i x) then f i x else 0) else (THE x. False))" + by (auto simp: lim_def convergent_eq_cauchy[symmetric]) + { fix x have "convergent (\i. if Cauchy (\i. f i x) then f i x else 0)" + by (cases "Cauchy (\i. f i x)") + (auto simp add: convergent_eq_cauchy[symmetric] convergent_def) } + note convergent = this + show ?thesis + unfolding * + apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const) + apply (rule borel_measurable_LIMSEQ) + apply (rule convergent_LIMSEQ_iff[THEN iffD1, OF convergent]) + apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const) + done +qed + +lemma borel_measurable_suminf: + fixes f :: "nat \ 'a \ real" + assumes f: "\i. f i \ borel_measurable M" + shows "(\x. suminf (\i. f i x)) \ borel_measurable M" + unfolding suminf_def sums_def[abs_def] lim_def[symmetric] + by (simp add: f borel_measurable_lim) + +end diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Wed Oct 10 13:30:50 2012 +0200 @@ -9,12 +9,6 @@ imports Measure_Space begin -lemma sums_def2: - "f sums x \ (\n. (\i\n. f i)) ----> x" - unfolding sums_def - apply (subst LIMSEQ_Suc_iff[symmetric]) - unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .. - text {* Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson. *} @@ -684,146 +678,6 @@ subsubsection {*Alternative instances of caratheodory*} -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 \ 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 \ 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) - moreover have "(\n. (\i=0.. (\i. f (disjointed A i))" - using f(1)[unfolded positive_def] dA - by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos) - from LIMSEQ_Suc[OF this] - have "(\n. (\i\n. f (disjointed A i))) ----> (\i. f (disjointed A i))" - unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost . - moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" - 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 \ 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) \ 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))" - using A - by (intro additive_sum[OF f, of _ A, symmetric]) - (auto intro: disjoint_family_on_mono[where B=UNIV]) - ultimately - have "(\i. f (A i)) sums f (\i. A i)" - unfolding sums_def2 by simp - from sums_unique[OF this] - show "(\i. f (A i)) = f (\i. A i)" by simp -qed - -lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: - assumes f: "positive M f" "additive M f" - 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 \ 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 \ 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 \ 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))" - using A by (auto simp: decseq_def intro!: f_mono) - have decseq: "decseq (\i. A i - (\i. A i))" - using A by (auto simp: decseq_def) - then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" - using A unfolding decseq_def by (auto intro!: f_mono Diff) - have "f (\x. A x) \ f (A 0)" - using A by (auto intro!: f_mono) - then have f_Int_fin: "f (\x. A x) \ \" - using A by auto - { fix i - have "f (A i - (\i. A i)) \ f (A i)" using A by (auto intro!: f_mono) - then have "f (A i - (\i. A i)) \ \" - using A by auto } - 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)) \ M" "(\i. A i - (\i. A i)) = {}" - using A by auto - qed - from INF_Lim_ereal[OF decseq_f this] - have "(INF n. f (A n - (\i. A i))) = 0" . - moreover have "(INF n. f (\i. A i)) = f (\i. A i)" - by auto - ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" - using A(4) f_fin f_Int_fin - by (subst INFI_ereal_add) (auto simp: decseq_f) - moreover { - fix n - have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" - using A by (subst f(2)[THEN additiveD]) auto - also have "(A n - (\i. A i)) \ (\i. A i) = A n" - by auto - finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } - ultimately have "(INF n. f (A n)) = f (\i. A i)" - by simp - with LIMSEQ_ereal_INFI[OF decseq_fA] - show "(\i. f (A i)) ----> f (\i. A i)" by simp -qed - -lemma positiveD1: "positive M f \ f {} = 0" 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\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\M. \x. f A = ereal x" - proof - 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] - from A have "(\i. f ((\i. A i) - A i)) ----> 0" - by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) - moreover - { fix i - have "f ((\i. A i) - A i) + f (A i) = f ((\i. A i) - A i \ A i)" - using A by (intro f(2)[THEN additiveD, symmetric]) auto - also have "(\i. A i) - A i \ A i = (\i. A i)" - by auto - finally have "f' (\i. A i) - f' (A i) = f' ((\i. A i) - A i)" - using A by (subst (asm) (1 2 3) f') auto - then have "f ((\i. A i) - A i) = ereal (f' (\i. A i) - f' (A i))" - using A f' by auto } - ultimately have "(\i. f' (\i. A i) - f' (A i)) ----> 0" - by (simp add: zero_ereal_def) - then have "(\i. f' (A i)) ----> f' (\i. A i)" - by (rule LIMSEQ_diff_approach_zero2[OF tendsto_const]) - then show "(\i. f (A i)) ----> f (\i. A i)" - using A by (subst (1 2) f') auto -qed - -lemma (in ring_of_sets) empty_continuous_imp_countably_additive: - 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 \ M \ f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -44,45 +44,45 @@ unfolding extensional_def by auto definition - "merge I x J y = (\i. if i \ I then x i else if i \ J then y i else undefined)" + "merge I J = (\(x, y) i. if i \ I then x i else if i \ J then y i else undefined)" lemma merge_apply[simp]: - "I \ J = {} \ i \ I \ merge I x J y i = x i" - "I \ J = {} \ i \ J \ merge I x J y i = y i" - "J \ I = {} \ i \ I \ merge I x J y i = x i" - "J \ I = {} \ i \ J \ merge I x J y i = y i" - "i \ I \ i \ J \ merge I x J y i = undefined" + "I \ J = {} \ i \ I \ merge I J (x, y) i = x i" + "I \ J = {} \ i \ J \ merge I J (x, y) i = y i" + "J \ I = {} \ i \ I \ merge I J (x, y) i = x i" + "J \ I = {} \ i \ J \ merge I J (x, y) i = y i" + "i \ I \ i \ J \ merge I J (x, y) i = undefined" unfolding merge_def by auto lemma merge_commute: - "I \ J = {} \ merge I x J y = merge J y I x" + "I \ J = {} \ merge I J (x, y) = merge J I (y, x)" by (auto simp: merge_def intro!: ext) lemma Pi_cancel_merge_range[simp]: - "I \ J = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" - "I \ J = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" - "J \ I = {} \ x \ Pi I (merge I A J B) \ x \ Pi I A" - "J \ I = {} \ x \ Pi I (merge J B I A) \ x \ Pi I A" + "I \ J = {} \ x \ Pi I (merge I J (A, B)) \ x \ Pi I A" + "I \ J = {} \ x \ Pi I (merge J I (B, A)) \ x \ Pi I A" + "J \ I = {} \ x \ Pi I (merge I J (A, B)) \ x \ Pi I A" + "J \ I = {} \ x \ Pi I (merge J I (B, A)) \ x \ Pi I A" by (auto simp: Pi_def) lemma Pi_cancel_merge[simp]: - "I \ J = {} \ merge I x J y \ Pi I B \ x \ Pi I B" - "J \ I = {} \ merge I x J y \ Pi I B \ x \ Pi I B" - "I \ J = {} \ merge I x J y \ Pi J B \ y \ Pi J B" - "J \ I = {} \ merge I x J y \ Pi J B \ y \ Pi J B" + "I \ J = {} \ merge I J (x, y) \ Pi I B \ x \ Pi I B" + "J \ I = {} \ merge I J (x, y) \ Pi I B \ x \ Pi I B" + "I \ J = {} \ merge I J (x, y) \ Pi J B \ y \ Pi J B" + "J \ I = {} \ merge I J (x, y) \ Pi J B \ y \ Pi J B" by (auto simp: Pi_def) -lemma extensional_merge[simp]: "merge I x J y \ extensional (I \ J)" +lemma extensional_merge[simp]: "merge I J (x, y) \ extensional (I \ J)" by (auto simp: extensional_def) lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" by (auto simp: restrict_def Pi_def) lemma restrict_merge[simp]: - "I \ J = {} \ restrict (merge I x J y) I = restrict x I" - "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" + "I \ J = {} \ restrict (merge I J (x, y)) I = restrict x I" + "I \ J = {} \ restrict (merge I J (x, y)) J = restrict y J" + "J \ I = {} \ restrict (merge I J (x, y)) I = restrict x I" + "J \ I = {} \ restrict (merge I J (x, y)) J = restrict y J" by (auto simp: restrict_def) lemma extensional_insert_undefined[intro, simp]: @@ -95,7 +95,7 @@ shows "a \ extensional (insert i I)" using assms unfolding extensional_def by auto -lemma merge_singleton[simp]: "i \ I \ merge I x {i} y = restrict (x(i := y i)) (insert i I)" +lemma merge_singleton[simp]: "i \ I \ merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)" unfolding merge_def by (auto simp: fun_eq_iff) lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" @@ -118,24 +118,24 @@ lemma restrict_vimage: assumes "I \ J = {}" - shows "(\x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \ Pi\<^isub>E J F) = Pi (I \ J) (merge I E J F)" + shows "(\x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \ Pi\<^isub>E J F) = Pi (I \ J) (merge I J (E, F))" using assms by (auto simp: restrict_Pi_cancel) lemma merge_vimage: assumes "I \ J = {}" - shows "(\(x,y). merge I x J y) -` Pi\<^isub>E (I \ J) E = Pi I E \ Pi J E" + shows "merge I J -` Pi\<^isub>E (I \ J) E = Pi I E \ Pi J E" 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) 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" + "merge I J (restrict x I, y) = merge I J (x, y)" + "merge I J (x, restrict y J) = merge I J (x, y)" unfolding merge_def by auto lemma merge_x_x_eq_restrict[simp]: - "merge I x J x = restrict x (I \ J)" + "merge I J (x, x) = restrict x (I \ J)" 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" @@ -504,8 +504,7 @@ 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)" + "(\(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)" @@ -519,19 +518,18 @@ qed (auto simp: space_pair_measure space_PiM) 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)" + "merge I J \ 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} = + then have "{\ \ space ?P. merge I 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 + by (auto simp: merge_def) 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) + finally show "{\ \ space ?P. merge I J \ i \ A} \ sets ?P" . +qed (auto simp: space_pair_measure space_PiM Pi_iff merge_def extensional_def) lemma measurable_restrict: assumes X: "\i. i \ I \ X i \ measurable N (M i)" @@ -587,44 +585,38 @@ qed qed -lemma (in product_sigma_finite) - assumes "finite 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 +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 emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\_. undefined} = 1" +proof - 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) ?\" + 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) + show "countably_additive (PiM {} M) ?\" + by (rule countably_additiveI_finite) + (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: ) qed (auto simp: prod_emb_def) - also have *: "(prod_emb {} M {} (\\<^isub>E i\{}. {})) = {\_. undefined}" + 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" . + finally show ?thesis + by simp +qed - interpret finite_measure "PiM {} M" - by default (simp add: space_PiM emeasure_eq) - case 1 show ?case .. +lemma PiM_empty: "PiM {} M = count_space {\_. undefined}" + by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def) - case 2 show ?case - using emeasure_eq * by simp -next +lemma (in product_sigma_finite) emeasure_PiM: + "finite I \ (\i. i\I \ A i \ sets (M i)) \ emeasure (PiM I M) (Pi\<^isub>E I A) = (\i\I. emeasure (M i) (A i))" +proof (induct I arbitrary: A rule: finite_induct) 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 - 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 ?P = "distr (Pi\<^isub>M I M \\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h" @@ -632,67 +624,73 @@ 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 = \" . + 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 M.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 + 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 by auto + qed (auto intro!: setprod_cong) + with insert show ?case + by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets_into_space) +qed (simp add: emeasure_PiM_empty) - 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) \ \" +lemma (in product_sigma_finite) sigma_finite: + assumes "finite I" + shows "sigma_finite_measure (PiM I M)" +proof - + interpret finite_product_sigma_finite M I by default fact + + from sigma_finite_pairs guess F :: "'i \ nat \ 'a set" .. + then have F: "\j. j \ I \ range (F j) \ sets (M j)" + "incseq (\k. \\<^isub>E j \ I. F j k)" + "(\k. \\<^isub>E j \ I. F j k) = space (Pi\<^isub>M I M)" + "\k. \j. j \ I \ emeasure (M j) (F j k) \ \" by blast+ - let ?F = "\k. \\<^isub>E j \ insert i I. F j k" + let ?F = "\k. \\<^isub>E j \ I. F j k" - case 1 show ?case + show ?thesis 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 + show "range ?F \ sets (Pi\<^isub>M I M)" using F(1) `finite I` by auto next - from F(3) show "(\i. ?F i) = space (Pi\<^isub>M (insert i I) M)" by simp + from F(3) show "(\i. ?F i) = space (Pi\<^isub>M 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 + from F `finite I` setprod_PInf[of I, OF emeasure_nonneg, of M] + show "emeasure (\\<^isub>M i\I. M i) (?F j) \ \" + by (subst emeasure_PiM) auto qed qed @@ -703,18 +701,6 @@ "(\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]: - "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 - 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)" @@ -734,7 +720,7 @@ lemma (in product_sigma_finite) distr_merge: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - 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" + shows "distr (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \ J) M) (merge I J) = Pi\<^isub>M (I \ J) M" (is "?D = ?P") proof - interpret I: finite_product_sigma_finite M I by default fact @@ -742,7 +728,7 @@ 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 "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default - let ?g = "\(x,y). merge I x J y" + let ?g = "merge I J" from IJ.sigma_finite_pairs obtain F where F: "\i. i\ I \ J \ range (F i) \ sets (M i)" @@ -765,7 +751,6 @@ 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 @@ -786,7 +771,7 @@ 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) + by (simp add: J.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 "\ = emeasure ?P (Pi\<^isub>E (I \ J) F)" @@ -800,16 +785,16 @@ 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))" + (\\<^isup>+ x. (\\<^isup>+ y. f (merge I J (x, y)) \(Pi\<^isub>M J M)) \(Pi\<^isub>M I M))" proof - interpret I: finite_product_sigma_finite M I by default fact interpret J: finite_product_sigma_finite M J by default fact interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default - 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) + have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)" + using measurable_comp[OF measurable_merge f] by (simp add: comp_def) show ?thesis apply (subst distr_merge[OF IJ, symmetric]) - apply (subst positive_integral_distr[OF measurable_merge f, OF IJ(1)]) + apply (subst positive_integral_distr[OF measurable_merge f]) apply (subst P.positive_integral_fst_measurable(2)[symmetric, OF P_borel]) apply simp done @@ -840,7 +825,7 @@ qed lemma (in product_sigma_finite) product_positive_integral_insert: - assumes [simp]: "finite I" "i \ I" + assumes I[simp]: "finite I" "i \ I" and f: "f \ borel_measurable (Pi\<^isub>M (insert i I) M)" shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\\<^isup>+ x. (\\<^isup>+ y. f (x(i := y)) \(M i)) \(Pi\<^isub>M I M))" proof - @@ -849,17 +834,17 @@ 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 positive_integral_cong, subst product_positive_integral_singleton) + unfolding product_positive_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f] + proof (rule positive_integral_cong, subst product_positive_integral_singleton[symmetric]) 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 space_PiM) - show "?f \ borel_measurable (M i)" unfolding f'_eq + let ?f = "\y. f (x(i := y))" + show "?f \ borel_measurable (M i)" 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 + show "(\\<^isup>+ y. f (merge I {i} (x, y)) \Pi\<^isub>M {i} M) = (\\<^isup>+ y. f (x(i := y i)) \Pi\<^isub>M {i} M)" + using x + by (auto intro!: positive_integral_cong arg_cong[where f=f] + simp add: space_PiM extensional_def) qed qed @@ -902,29 +887,54 @@ lemma (in product_sigma_finite) product_integral_fold: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" and f: "integrable (Pi\<^isub>M (I \ J) M) f" - shows "integral\<^isup>L (Pi\<^isub>M (I \ J) M) f = (\x. (\y. f (merge I x J y) \Pi\<^isub>M J M) \Pi\<^isub>M I M)" + shows "integral\<^isup>L (Pi\<^isub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^isub>M J M) \Pi\<^isub>M I M)" proof - interpret I: finite_product_sigma_finite M I by default fact interpret J: finite_product_sigma_finite M J by default fact 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 "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default - let ?M = "\(x, y). merge I x J y" + let ?M = "merge I J" 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 P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^isub>M I M \\<^isub>M Pi\<^isub>M J M)" + using measurable_comp[OF measurable_merge 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) + by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) show ?thesis apply (subst distr_merge[symmetric, OF IJ fin]) - apply (subst integral_distr[OF measurable_merge[OF IJ] f_borel]) + apply (subst integral_distr[OF measurable_merge f_borel]) apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int]) apply simp done qed +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) ((\y. merge I J (x, y)) -` 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) ((\y. merge I J (x, y)) -` 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: "merge I J -` 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 A]) + apply (subst J.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 (in product_sigma_finite) product_integral_insert: assumes I: "finite I" "i \ I" and f: "integrable (Pi\<^isub>M (insert i I) M) f" @@ -932,7 +942,7 @@ proof - 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)" + also have "\ = (\x. (\y. f (merge I {i} (x,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]) @@ -942,7 +952,7 @@ 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)" + from x I show "(\ y. f (merge I {i} (x,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 . @@ -980,30 +990,83 @@ lemma (in product_sigma_finite) product_integral_setprod: fixes f :: "'i \ 'a \ real" - assumes "finite I" "I \ {}" and integrable: "\i. i \ I \ integrable (M i) (f i)" + assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "(\x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>L (M i) (f i))" -using assms proof (induct rule: finite_ne_induct) - case (singleton i) - then show ?case by (simp add: product_integral_singleton integrable_def) +using assms proof induct + case empty + interpret finite_measure "Pi\<^isub>M {} M" + by rule (simp add: space_PiM) + show ?case by (simp add: space_PiM measure_def) next case (insert i I) then have iI: "finite (insert i I)" by auto then have prod: "\J. J \ insert i I \ integrable (Pi\<^isub>M J M) (\x. (\i\J. f i (x i)))" - by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) + by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) interpret I: finite_product_sigma_finite M I by default fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using `i \ I` by (auto intro!: setprod_cong) show ?case - unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] + unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI) qed +lemma sets_Collect_single: + "i \ I \ A \ sets (M i) \ { x \ space (Pi\<^isub>M I M). x i \ A } \ sets (Pi\<^isub>M I M)" + unfolding sets_PiM_single + by (auto intro!: sigma_sets.Basic exI[of _ i] exI[of _ A]) (auto simp: space_PiM) + +lemma sigma_prod_algebra_sigma_eq_infinite: + fixes E :: "'i \ 'a set set" + assumes 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 == {{f\\\<^isub>E i\I. space (M i). f i \ A} | i A. i \ I \ A \ 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)" + then have "(\x. x i) \ measurable ?P (sigma (space (M i)) (E i))" + apply (subst measurable_iff_measure_of) + apply (simp_all add: P_closed) + using E_closed + apply (force simp: subset_eq space_PiM) + apply (force simp: subset_eq space_PiM) + apply (auto simp: P_def intro!: sigma_sets.Basic exI[of _ i]) + apply (rule_tac x=Aa in exI) + apply (auto simp: space_PiM) + done + 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)" + unfolding P_def space_PiM[symmetric] + by (intro sigma_sets_subset) (auto simp: E_generates sets_Collect_single) +qed + +lemma bchoice_iff: "(\a\A. \b. P a b) \ (\f. \a\A. P a (f a))" + by metis + lemma sigma_prod_algebra_sigma_eq: - fixes E :: "'i \ 'a set set" + fixes E :: "'i \ 'a set set" and S :: "'i \ nat \ 'a 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)" + assumes 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)" @@ -1011,6 +1074,9 @@ shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P" proof let ?P = "sigma (space (Pi\<^isub>M I M)) P" + from `finite I`[THEN ex_bij_betw_finite_nat] guess T .. + then have T: "\i. i \ I \ T i < card I" "\i. i\I \ the_inv_into I T (T i) = i" + by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f) 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))" @@ -1033,15 +1099,18 @@ 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 "\ = (\xs\{xs. length xs = card I}. \\<^isub>E j\I. if i = j then A else S j (xs ! T j))" + using T + apply (auto simp: Pi_iff bchoice_iff) + apply (rule_tac x="map (\n. f (the_inv_into I T n)) [0.. \ 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" + fix xs show "(\\<^isub>E j\I. if i = j then A else S j (xs ! T j)) \ 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"]) + (auto simp: P_def subset_eq intro!: exI[of _ "\j. if i = j then A else S j (xs ! T j)"]) qed finally show "(\x. x i) -` A \ space ?P \ sets ?P" using P_closed by simp diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Independent_Family.thy Wed Oct 10 13:30:50 2012 +0200 @@ -23,13 +23,6 @@ qed auto definition (in prob_space) - "indep_events A I \ (A`I \ events) \ - (\J\I. J \ {} \ finite J \ prob (\j\J. A j) = (\j\J. prob (A j)))" - -definition (in prob_space) - "indep_event A B \ indep_events (bool_case A B) UNIV" - -definition (in prob_space) "indep_sets F I \ (\i\I. F i \ events) \ (\J\I. J \ {} \ finite J \ (\A\Pi J F. prob (\j\J. A j) = (\j\J. prob (A j))))" @@ -37,22 +30,27 @@ "indep_set A B \ indep_sets (bool_case A B) UNIV" definition (in prob_space) - "indep_vars M' X I \ - (\i\I. random_variable (M' i) (X i)) \ - indep_sets (\i. sigma_sets (space M) { X i -` A \ space M | A. A \ sets (M' i)}) I" + indep_events_def_alt: "indep_events A I \ indep_sets (\i. {A i}) I" + +lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" + by auto + +lemma (in prob_space) indep_events_def: + "indep_events A I \ (A`I \ events) \ + (\J\I. J \ {} \ finite J \ prob (\j\J. A j) = (\j\J. prob (A j)))" + unfolding indep_events_def_alt indep_sets_def + apply (simp add: Ball_def Pi_iff image_subset_iff_funcset) + apply (intro conj_cong refl arg_cong[where f=All] ext imp_cong) + apply auto + done definition (in prob_space) - "indep_var Ma A Mb B \ indep_vars (bool_case Ma Mb) (bool_case A B) UNIV" + "indep_event A B \ indep_events (bool_case A B) UNIV" 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+ -lemma (in prob_space) indep_sets_singleton_iff_indep_events: - "indep_sets (\i. {F i}) I \ indep_events F I" - unfolding indep_sets_def indep_events_def - by (simp, intro conj_cong ball_cong all_cong imp_cong) (auto simp: Pi_iff) - lemma (in prob_space) indep_events_finite_index_events: "indep_events F I \ (\J\I. J \ {} \ finite J \ indep_events F J)" by (auto simp: indep_events_def) @@ -86,6 +84,15 @@ using indep by (auto simp: indep_sets_def) qed +lemma (in prob_space) indep_sets_mono: + assumes indep: "indep_sets F I" + assumes mono: "J \ I" "\i. i\J \ G i \ F i" + shows "indep_sets G J" + apply (rule indep_sets_mono_sets) + apply (rule indep_sets_mono_index) + apply (fact +) + done + lemma (in prob_space) indep_setsI: assumes "\i. i \ I \ F i \ events" and "\A J. J \ {} \ J \ I \ finite J \ (\j\J. A j \ F j) \ prob (\j\J. A j) = (\j\J. prob (A j))" @@ -117,16 +124,6 @@ using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev by (simp add: ac_simps UNIV_bool) -lemma (in prob_space) indep_var_eq: - "indep_var S X T Y \ - (random_variable S X \ random_variable T Y) \ - indep_set - (sigma_sets (space M) { X -` A \ space M | A. A \ sets S}) - (sigma_sets (space M) { Y -` A \ space M | A. A \ sets T})" - unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool - by (intro arg_cong2[where f="op \"] arg_cong2[where f=indep_sets] ext) - (auto split: bool.split) - lemma (in prob_space) assumes indep: "indep_set A B" shows indep_setD_ev1: "A \ events" @@ -326,6 +323,36 @@ by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic) qed +definition (in prob_space) + indep_vars_def2: "indep_vars M' X I \ + (\i\I. random_variable (M' i) (X i)) \ + indep_sets (\i. { X i -` A \ space M | A. A \ sets (M' i)}) I" + +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_vars_def: + "indep_vars M' X I \ + (\i\I. random_variable (M' i) (X i)) \ + indep_sets (\i. sigma_sets (space M) { X i -` A \ space M | A. A \ sets (M' i)}) I" + unfolding indep_vars_def2 + apply (rule conj_cong[OF refl]) + apply (rule indep_sets_sigma_sets_iff[symmetric]) + apply (auto simp: Int_stable_def) + apply (rule_tac x="A \ Aa" in exI) + apply auto + done + +lemma (in prob_space) indep_var_eq: + "indep_var S X T Y \ + (random_variable S X \ random_variable T Y) \ + indep_set + (sigma_sets (space M) { X -` A \ space M | A. A \ sets S}) + (sigma_sets (space M) { Y -` A \ space M | A. A \ sets T})" + unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool + by (intro arg_cong2[where f="op \"] arg_cong2[where f=indep_sets] ext) + (auto split: bool.split) + lemma (in prob_space) indep_sets2_eq: "indep_set A B \ A \ events \ B \ events \ (\a\A. \b\B. prob (a \ b) = prob a * prob b)" unfolding indep_set_def @@ -468,27 +495,25 @@ by (simp cong: indep_sets_cong) qed -definition (in prob_space) terminal_events where - "terminal_events A = (\n. sigma_sets (space M) (UNION {n..} A))" +definition (in prob_space) tail_events where + "tail_events A = (\n. sigma_sets (space M) (UNION {n..} A))" -lemma (in prob_space) terminal_events_sets: - assumes A: "\i. A i \ events" - assumes "\i::nat. sigma_algebra (space M) (A i)" - assumes X: "X \ terminal_events A" - shows "X \ events" -proof - +lemma (in prob_space) tail_events_sets: + assumes A: "\i::nat. A i \ events" + shows "tail_events A \ events" +proof + fix X assume X: "X \ tail_events A" let ?A = "(\n. sigma_sets (space M) (UNION {n..} A))" - 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 X have "\n::nat. X \ sigma_sets (space M) (UNION {n..} A)" by (auto simp: tail_events_def) from this[of 0] have "X \ sigma_sets (space M) (UNION UNIV A)" by simp then show "X \ events" by induct (insert A, auto) qed -lemma (in prob_space) sigma_algebra_terminal_events: +lemma (in prob_space) sigma_algebra_tail_events: assumes "\i::nat. sigma_algebra (space M) (A i)" - shows "sigma_algebra (space M) (terminal_events A)" - unfolding terminal_events_def + shows "sigma_algebra (space M) (tail_events A)" + unfolding tail_events_def proof (simp add: sigma_algebra_iff2, safe) let ?A = "(\n. sigma_sets (space M) (UNION {n..} A))" interpret A: sigma_algebra "space M" "A i" for i by fact @@ -505,20 +530,22 @@ 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 M) (A i)" assumes indep: "indep_sets A UNIV" - and X: "X \ terminal_events A" + and X: "X \ tail_events A" shows "prob X = 0 \ prob X = 1" proof - + have A: "\i. A i \ events" + using indep unfolding indep_sets_def by simp + 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 + interpret T: sigma_algebra "space M" "tail_events A" + by (rule sigma_algebra_tail_events) fact have "X \ space M" using T.space_closed X by auto have X_in: "X \ events" - by (rule terminal_events_sets) fact+ + using tail_events_sets A X by auto interpret D: dynkin_system "space M" ?D proof (rule dynkin_systemI) @@ -583,7 +610,7 @@ 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)" - using X unfolding terminal_events_def by simp + using X unfolding tail_events_def by simp from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D show "D \ events \ prob (X \ D) = prob X * prob D" by (auto simp add: ac_simps) @@ -591,12 +618,12 @@ then have "(\n. sigma_sets (space M) (\m\{..n}. A m)) \ ?D" (is "?A \ _") by auto - note `X \ terminal_events A` + note `X \ tail_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 } + then have "tail_events A \ sigma_sets (space M) ?A" + unfolding tail_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)" @@ -628,34 +655,33 @@ lemma (in prob_space) borel_0_1_law: fixes F :: "nat \ 'a set" - assumes F: "range F \ events" "indep_events F UNIV" + assumes F2: "indep_events F UNIV" shows "prob (\n. \m\{n..}. F m) = 0 \ prob (\n. \m\{n..}. F m) = 1" proof (rule kolmogorov_0_1_law[of "\i. sigma_sets (space M) { F i }"]) - show "\i. sigma_sets (space M) {F i} \ events" - using F(1) sets_into_space - by (subst sigma_sets_singleton) auto + have F1: "range F \ events" + using F2 by (simp add: indep_events_def subset_eq) { 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 + using sigma_algebra_sigma_sets[of "{F i}" "space M"] F1 sets_into_space by auto } show "indep_sets (\i. sigma_sets (space M) {F i}) UNIV" proof (rule indep_sets_sigma) show "indep_sets (\i. {F i}) UNIV" - unfolding indep_sets_singleton_iff_indep_events by fact + unfolding indep_events_def_alt[symmetric] by fact fix i show "Int_stable {F i}" unfolding Int_stable_def by simp qed let ?Q = "\n. \i\{n..}. F i" - show "(\n. \m\{n..}. F m) \ terminal_events (\i. sigma_sets (space M) {F i})" - unfolding terminal_events_def + show "(\n. \m\{n..}. F m) \ tail_events (\i. sigma_sets (space M) {F i})" + unfolding tail_events_def proof fix j 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] + using order_trans[OF F1 space_closed] 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 "\ \ sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" - using order_trans[OF F(1) space_closed] + using order_trans[OF F1 space_closed] by (safe intro!: S.countable_INT S.countable_UN) (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})" @@ -872,7 +898,7 @@ 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')" + show "range ?A \ prod_algebra I M'" "(\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 @@ -1022,13 +1048,13 @@ 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) \ \" + show "range F \ ?E" "(\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) + by (simp add: M2.emeasure_pair_measure_Times) also have "\ = emeasure M (A \ B)" using A B emeasure by auto finally show "emeasure ?P X = emeasure M X" @@ -1136,7 +1162,7 @@ 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) + using ab by (simp add: Y.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) @@ -1144,4 +1170,12 @@ qed qed +lemma (in prob_space) distributed_joint_indep: + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes X: "distributed M S X Px" and Y: "distributed M T Y Py" + assumes indep: "indep_var S X T Y" + shows "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) (\(x, y). Px x * Py y)" + using indep_var_distribution_eq[of S X T Y] indep + by (intro distributed_joint_indep'[OF S T X Y]) auto + end diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Infinite_Product_Measure.thy --- a/src/HOL/Probability/Infinite_Product_Measure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -8,41 +8,16 @@ imports Probability_Measure Caratheodory begin -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 lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" unfolding restrict_def by (simp add: fun_eq_iff) -lemma split_merge: "P (merge I x J y i) \ (i \ I \ P (x i)) \ (i \ J - I \ P (y i)) \ (i \ I \ J \ P undefined)" +lemma split_merge: "P (merge I J (x,y) i) \ (i \ I \ P (x i)) \ (i \ J - I \ P (y i)) \ (i \ I \ J \ P undefined)" unfolding merge_def by auto -lemma extensional_merge_sub: "I \ J \ K \ merge I x J y \ extensional K" +lemma extensional_merge_sub: "I \ J \ K \ merge I J (x, y) \ extensional K" unfolding merge_def extensional_def by auto lemma injective_vimage_restrict: @@ -57,7 +32,7 @@ show "x \ A \ x \ B" proof cases assume x: "x \ (\\<^isub>E i\J. S i)" - have "x \ A \ merge J x (I - J) y \ (\x. restrict x J) -` A \ (\\<^isub>E i\I. S i)" + have "x \ A \ merge J (I - J) (x,y) \ (\x. restrict x J) -` A \ (\\<^isub>E i\I. S i)" using y x `J \ I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge) then show "x \ A \ x \ B" using y x `J \ I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge) @@ -112,7 +87,7 @@ 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) = ?\" + show "range ?F \ ?J" "(\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) @@ -156,7 +131,7 @@ 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 + 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))" @@ -199,15 +174,24 @@ qed 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) + "sets (PiM I M) = sigma_sets (\\<^isub>E i\I. space (M i)) generator" +proof cases + assume "I = {}" then show ?thesis + unfolding generator_def + by (auto simp: sets_PiM_empty sigma_sets_empty_eq cong: conj_cong) +next + assume "I \ {}" + show ?thesis + 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) +qed + lemma (in product_prob_space) generatorI: "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ A \ generator" @@ -255,23 +239,19 @@ qed lemma (in product_prob_space) merge_sets: - 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 - - 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]) -qed + assumes "J \ K = {}" and A: "A \ sets (Pi\<^isub>M (J \ K) M)" and x: "x \ space (Pi\<^isub>M J M)" + shows "(\y. merge J K (x,y)) -` A \ space (Pi\<^isub>M K M) \ sets (Pi\<^isub>M K M)" + by (rule measurable_sets[OF _ A] measurable_compose[OF measurable_Pair measurable_merge] + measurable_const x measurable_ident)+ lemma (in product_prob_space) merge_emb: assumes "K \ I" "J \ I" and y: "y \ space (Pi\<^isub>M J M)" - shows "(merge J y (I - J) -` emb I K X \ space (Pi\<^isub>M I M)) = - emb I (K - J) (merge J y (K - J) -` emb (J \ K) K X \ space (Pi\<^isub>M (K - J) M))" + shows "((\x. merge J (I - J) (y, x)) -` emb I K X \ space (Pi\<^isub>M I M)) = + emb I (K - J) ((\x. merge J (K - J) (y, x)) -` emb (J \ K) K X \ space (Pi\<^isub>M (K - J) M))" proof - - have [simp]: "\x J K L. merge J y K (restrict x L) = merge J y (K \ L) x" + have [simp]: "\x J K L. merge J K (y, restrict x L) = merge J (K \ L) (y, x)" by (auto simp: restrict_def merge_def) - have [simp]: "\x J K L. restrict (merge J y K x) L = merge (J \ L) y (K \ L) x" + have [simp]: "\x J K L. restrict (merge J K (y, x)) L = merge (J \ L) (K \ L) (y, x)" by (auto simp: restrict_def merge_def) have [simp]: "(I - J) \ K = K - J" using `K \ I` `J \ I` by auto have [simp]: "(K - J) \ (K \ J) = K - J" by auto @@ -356,16 +336,16 @@ "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)" + let ?M = "\y. (\x. merge J (K - J) (y, x)) -` emb (J \ K) K X \ space (Pi\<^isub>M (K - J) M)" { fix y assume y: "y \ space (Pi\<^isub>M J M)" note * = merge_emb[OF `K \ I` `J \ I` y, of X] moreover 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)) \ ?G" + have ***: "((\x. merge J (I - J) (y, x)) -` 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)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)" + have "\G ((\x. merge J (I - J) (y, x)) -` 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 @@ -379,7 +359,7 @@ using K J by simp 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)" + also have "\ = (\\<^isup>+ y. \G ((\x. merge J (I - J) (y, x)) -` 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 positive_integral_cong) fix x assume x: "x \ space (Pi\<^isub>M J M)" @@ -395,7 +375,7 @@ by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) } note le_1 = this - let ?q = "\y. \G (merge J y (I - J) -` Z \ space (Pi\<^isub>M I M))" + let ?q = "\y. \G ((\x. merge J (I - J) (y,x)) -` 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 emeasure_fold_measurable cong: measurable_cong) @@ -441,7 +421,7 @@ using \G_spec[of "J 0" "A 0" "X 0"] J A_eq by (auto intro!: INF_lower2[of 0] J.measure_le_1) - let ?M = "\K Z y. merge K y (I - K) -` Z \ space (Pi\<^isub>M I M)" + let ?M = "\K Z y. (\x. merge K (I - K) (y, x)) -` Z \ space (Pi\<^isub>M I M)" { 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 @@ -549,9 +529,9 @@ from Ex_w[OF this `?D \ {}`] J[of "Suc k"] obtain w' where w': "w' \ space (Pi\<^isub>M ?D M)" "\n. ?a / 2 ^ (Suc k + 1) \ \G (?M ?D (?M (J k) (A n) (w k)) w')" by auto - let ?w = "merge (J k) (w k) ?D w'" - have [simp]: "\x. merge (J k) (w k) (I - J k) (merge ?D w' (I - ?D) x) = - merge (J (Suc k)) ?w (I - (J (Suc k))) x" + let ?w = "merge (J k) ?D (w k, w')" + have [simp]: "\x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) = + merge (J (Suc k)) (I - (J (Suc k))) (?w, x)" using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"] 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" @@ -577,7 +557,7 @@ using positive_\G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) then obtain x where "x \ ?M (J k) (A k) (w k)" by auto - then have "merge (J k) (w k) (I - J k) x \ A k" by auto + then have "merge (J k) (I - J k) (w 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 space_PiM) @@ -657,7 +637,7 @@ 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) + using \ unfolding sets_PiM_generator by (auto simp: measure_space_def) qed qed diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Information.thy Wed Oct 10 13:30:50 2012 +0200 @@ -22,104 +22,6 @@ "(\x\A \ B. f x) = (\x\A. setsum (\y. f (x, y)) B)" unfolding setsum_cartesian_product by simp -section "Convex theory" - -lemma log_setsum: - assumes "finite s" "s \ {}" - assumes "b > 1" - assumes "(\ i \ s. a i) = 1" - assumes "\ i. i \ s \ a i \ 0" - assumes "\ i. i \ s \ y i \ {0 <..}" - shows "log b (\ i \ s. a i * y i) \ (\ i \ s. a i * log b (y i))" -proof - - have "convex_on {0 <..} (\ x. - log b x)" - by (rule minus_log_convex[OF `b > 1`]) - hence "- log b (\ i \ s. a i * y i) \ (\ i \ s. a i * - log b (y i))" - using convex_on_setsum[of _ _ "\ x. - log b x"] assms pos_is_convex by fastforce - thus ?thesis by (auto simp add:setsum_negf le_imp_neg_le) -qed - -lemma log_setsum': - assumes "finite s" "s \ {}" - assumes "b > 1" - assumes "(\ i \ s. a i) = 1" - assumes pos: "\ i. i \ s \ 0 \ a i" - "\ i. \ i \ s ; 0 < a i \ \ 0 < y i" - shows "log b (\ i \ s. a i * y i) \ (\ i \ s. a i * log b (y i))" -proof - - have "\y. (\ i \ s - {i. a i = 0}. a i * y i) = (\ i \ s. a i * y i)" - using assms by (auto intro!: setsum_mono_zero_cong_left) - moreover have "log b (\ i \ s - {i. a i = 0}. a i * y i) \ (\ i \ s - {i. a i = 0}. a i * log b (y i))" - proof (rule log_setsum) - have "setsum a (s - {i. a i = 0}) = setsum a s" - using assms(1) by (rule setsum_mono_zero_cong_left) auto - thus sum_1: "setsum a (s - {i. a i = 0}) = 1" - "finite (s - {i. a i = 0})" using assms by simp_all - - show "s - {i. a i = 0} \ {}" - proof - assume *: "s - {i. a i = 0} = {}" - hence "setsum a (s - {i. a i = 0}) = 0" by (simp add: * setsum_empty) - with sum_1 show False by simp - qed - - fix i assume "i \ s - {i. a i = 0}" - hence "i \ s" "a i \ 0" by simp_all - thus "0 \ a i" "y i \ {0<..}" using pos[of i] by auto - qed fact+ - ultimately show ?thesis by simp -qed - -lemma log_setsum_divide: - assumes "finite S" and "S \ {}" and "1 < b" - assumes "(\x\S. g x) = 1" - assumes pos: "\x. x \ S \ g x \ 0" "\x. x \ S \ f x \ 0" - assumes g_pos: "\x. \ x \ S ; 0 < g x \ \ 0 < f x" - shows "- (\x\S. g x * log b (g x / f x)) \ log b (\x\S. f x)" -proof - - have log_mono: "\x y. 0 < x \ x \ y \ log b x \ log b y" - using `1 < b` by (subst log_le_cancel_iff) auto - - have "- (\x\S. g x * log b (g x / f x)) = (\x\S. g x * log b (f x / g x))" - proof (unfold setsum_negf[symmetric], rule setsum_cong) - fix x assume x: "x \ S" - show "- (g x * log b (g x / f x)) = g x * log b (f x / g x)" - proof (cases "g x = 0") - case False - with pos[OF x] g_pos[OF x] have "0 < f x" "0 < g x" by simp_all - thus ?thesis using `1 < b` by (simp add: log_divide field_simps) - qed simp - qed rule - also have "... \ log b (\x\S. g x * (f x / g x))" - proof (rule log_setsum') - fix x assume x: "x \ S" "0 < g x" - with g_pos[OF x] show "0 < f x / g x" by (safe intro!: divide_pos_pos) - qed fact+ - also have "... = log b (\x\S - {x. g x = 0}. f x)" using `finite S` - by (auto intro!: setsum_mono_zero_cong_right arg_cong[where f="log b"] - split: split_if_asm) - also have "... \ log b (\x\S. f x)" - proof (rule log_mono) - have "0 = (\x\S - {x. g x = 0}. 0)" by simp - also have "... < (\x\S - {x. g x = 0}. f x)" (is "_ < ?sum") - proof (rule setsum_strict_mono) - show "finite (S - {x. g x = 0})" using `finite S` by simp - show "S - {x. g x = 0} \ {}" - proof - assume "S - {x. g x = 0} = {}" - hence "(\x\S. g x) = 0" by (subst setsum_0') auto - with `(\x\S. g x) = 1` show False by simp - qed - fix x assume "x \ S - {x. g x = 0}" - thus "0 < f x" using g_pos[of x] pos(1)[of x] by auto - qed - finally show "0 < ?sum" . - show "(\x\S - {x. g x = 0}. f x) \ (\x\S. f x)" - using `finite S` pos by (auto intro!: setsum_mono2) - qed - finally show ?thesis . -qed - lemma split_pairs: "((A, B) = X) \ (fst X = A \ snd X = B)" and "(X = (A, B)) \ (fst X = A \ snd X = B)" by auto @@ -194,7 +96,7 @@ 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` + using f 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 @@ -437,6 +339,115 @@ finally show ?thesis . qed +subsection {* Finite Entropy *} + +definition (in information_space) + "finite_entropy S X f \ distributed M S X f \ integrable S (\x. f x * log b (f x))" + +lemma (in information_space) finite_entropy_simple_function: + assumes X: "simple_function M X" + shows "finite_entropy (count_space (X`space M)) X (\a. measure M {x \ space M. X x = a})" + unfolding finite_entropy_def +proof + have [simp]: "finite (X ` space M)" + using X by (auto simp: simple_function_def) + then show "integrable (count_space (X ` space M)) + (\x. prob {xa \ space M. X xa = x} * log b (prob {xa \ space M. X xa = x}))" + by (rule integrable_count_space) + have d: "distributed M (count_space (X ` space M)) X (\x. ereal (if x \ X`space M then prob {xa \ space M. X xa = x} else 0))" + by (rule distributed_simple_function_superset[OF X]) (auto intro!: arg_cong[where f=prob]) + show "distributed M (count_space (X ` space M)) X (\x. ereal (prob {xa \ space M. X xa = x}))" + by (rule distributed_cong_density[THEN iffD1, OF _ _ _ d]) auto +qed + +lemma distributed_transform_AE: + assumes T: "T \ measurable P Q" "absolutely_continuous Q (distr P Q T)" + assumes g: "distributed M Q Y g" + shows "AE x in P. 0 \ g (T x)" + using g + apply (subst AE_distr_iff[symmetric, OF T(1)]) + apply (simp add: distributed_borel_measurable) + apply (rule absolutely_continuous_AE[OF _ T(2)]) + apply simp + apply (simp add: distributed_AE) + done + +lemma ac_fst: + assumes "sigma_finite_measure T" + shows "absolutely_continuous S (distr (S \\<^isub>M T) S fst)" +proof - + interpret sigma_finite_measure T by fact + { fix A assume "A \ sets S" "emeasure S A = 0" + moreover then have "fst -` A \ space (S \\<^isub>M T) = A \ space T" + by (auto simp: space_pair_measure dest!: sets_into_space) + ultimately have "emeasure (S \\<^isub>M T) (fst -` A \ space (S \\<^isub>M T)) = 0" + by (simp add: emeasure_pair_measure_Times) } + then show ?thesis + unfolding absolutely_continuous_def + apply (auto simp: null_sets_distr_iff) + apply (auto simp: null_sets_def intro!: measurable_sets) + done +qed + +lemma ac_snd: + assumes "sigma_finite_measure T" + shows "absolutely_continuous T (distr (S \\<^isub>M T) T snd)" +proof - + interpret sigma_finite_measure T by fact + { fix A assume "A \ sets T" "emeasure T A = 0" + moreover then have "snd -` A \ space (S \\<^isub>M T) = space S \ A" + by (auto simp: space_pair_measure dest!: sets_into_space) + ultimately have "emeasure (S \\<^isub>M T) (snd -` A \ space (S \\<^isub>M T)) = 0" + by (simp add: emeasure_pair_measure_Times) } + then show ?thesis + unfolding absolutely_continuous_def + apply (auto simp: null_sets_distr_iff) + apply (auto simp: null_sets_def intro!: measurable_sets) + done +qed + +lemma distributed_integrable: + "distributed M N X f \ g \ borel_measurable N \ + integrable N (\x. f x * g x) \ integrable M (\x. g (X x))" + by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable + distributed_distr_eq_density[symmetric] integral_density[symmetric] integrable_distr_eq) + +lemma distributed_transform_integrable: + assumes Px: "distributed M N X Px" + assumes "distributed M P Y Py" + assumes Y: "Y = (\x. T (X x))" and T: "T \ measurable N P" and f: "f \ borel_measurable P" + shows "integrable P (\x. Py x * f x) \ integrable N (\x. Px x * f (T x))" +proof - + have "integrable P (\x. Py x * f x) \ integrable M (\x. f (Y x))" + by (rule distributed_integrable) fact+ + also have "\ \ integrable M (\x. f (T (X x)))" + using Y by simp + also have "\ \ integrable N (\x. Px x * f (T x))" + using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def) + finally show ?thesis . +qed + +lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" + using integrable_cong_AE by blast + +lemma (in information_space) finite_entropy_integrable: + "finite_entropy S X Px \ integrable S (\x. Px x * log b (Px x))" + unfolding finite_entropy_def by auto + +lemma (in information_space) finite_entropy_distributed: + "finite_entropy S X Px \ distributed M S X Px" + unfolding finite_entropy_def by auto + +lemma (in information_space) finite_entropy_integrable_transform: + assumes Fx: "finite_entropy S X Px" + assumes Fy: "distributed M T Y Py" + and "X = (\x. f (Y x))" + and "f \ measurable T S" + shows "integrable T (\x. Py x * log b (Px (f x)))" + using assms unfolding finite_entropy_def + using distributed_transform_integrable[of M T Y Py S X Px f "\x. log b (Px x)"] + by (auto intro: distributed_real_measurable) + subsection {* Mutual Information *} definition (in prob_space) @@ -510,6 +521,120 @@ lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes Fx: "finite_entropy S X Px" and Fy: "finite_entropy T Y Py" + assumes Fxy: "finite_entropy (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': "0 \ mutual_information b S T X Y" +proof - + have Px: "distributed M S X Px" + using Fx by (auto simp: finite_entropy_def) + have Py: "distributed M T Y Py" + using Fy by (auto simp: finite_entropy_def) + have Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + using Fxy by (auto simp: finite_entropy_def) + + have X: "random_variable S X" + using Px by (auto simp: distributed_def finite_entropy_def) + have Y: "random_variable T Y" + using Py by (auto simp: distributed_def finite_entropy_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) + + have X: "X = fst \ (\x. (X x, Y x))" and Y: "Y = snd \ (\x. (X x, Y x))" + by auto + + have "integrable (S \\<^isub>M T) (\x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))" + using finite_entropy_integrable[OF Fxy] + using finite_entropy_integrable_transform[OF Fx Pxy, of fst] + using finite_entropy_integrable_transform[OF Fy Pxy, of snd] + by simp + moreover have "f \ borel_measurable (S \\<^isub>M T)" + unfolding f_def using Px Py Pxy + by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'' + intro!: borel_measurable_times borel_measurable_log borel_measurable_divide) + ultimately have int: "integrable (S \\<^isub>M T) f" + apply (rule integrable_cong_AE_imp) + using + distributed_transform_AE[OF measurable_fst ac_fst, of T, OF T Px] + distributed_transform_AE[OF measurable_snd ac_snd, of _ _ _ _ S, OF T Py] + subdensity_real[OF measurable_fst Pxy Px X] + subdensity_real[OF measurable_snd Pxy Py Y] + distributed_real_AE[OF Pxy] + by eventually_elim + (auto simp: f_def log_divide_eq log_mult_eq field_simps zero_less_mult_iff mult_nonneg_nonneg) + + 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) + 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" @@ -664,84 +789,209 @@ entropy_Pow ("\'(_')") where "\(X) \ entropy b (count_space (X`space M)) X" -lemma (in information_space) entropy_distr: +lemma (in prob_space) distributed_RN_deriv: + assumes X: "distributed M S X Px" + shows "AE x in S. RN_deriv S (density S Px) x = Px x" +proof - + note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X] + interpret X: prob_space "distr M S X" + using D(1) by (rule prob_space_distr) + + have sf: "sigma_finite_measure (distr M S X)" by default + show ?thesis + using D + apply (subst eq_commute) + apply (intro RN_deriv_unique_sigma_finite) + apply (auto intro: divide_nonneg_nonneg measure_nonneg + simp: distributed_distr_eq_density[symmetric, OF X] sf) + done +qed + +lemma (in information_space) 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)" + assumes X: "distributed M MX X f" + shows entropy_distr: "entropy b MX X = - (\x. f x * log b (f x) \MX)" (is ?eq) +proof - + note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X] + note ae = distributed_RN_deriv[OF X] + + have ae_eq: "AE x in distr M MX X. log b (real (RN_deriv MX (distr M MX X) x)) = + log b (f x)" + unfolding distributed_distr_eq_density[OF X] + apply (subst AE_density) + using D apply simp + using ae apply eventually_elim + apply auto + done + + have int_eq: "- (\ x. log b (f x) \distr M MX X) = - (\ x. f x * log b (f x) \MX)" + unfolding distributed_distr_eq_density[OF X] + using D + by (subst integral_density) + (auto simp: borel_measurable_ereal_iff) + + show ?eq + unfolding entropy_def KL_divergence_def entropy_density_def comp_def + apply (subst integral_cong_AE) + apply (rule ae_eq) + apply (rule int_eq) + done +qed + +lemma (in prob_space) distributed_imp_emeasure_nonzero: + assumes X: "distributed M MX X Px" + shows "emeasure MX {x \ space MX. Px x \ 0} \ 0" +proof + note Px = distributed_borel_measurable[OF X] distributed_AE[OF X] + interpret X: prob_space "distr M MX X" + using distributed_measurable[OF X] by (rule prob_space_distr) + + assume "emeasure MX {x \ space MX. Px x \ 0} = 0" + with Px have "AE x in MX. Px x = 0" + by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ereal_iff) + moreover + from X.emeasure_space_1 have "(\\<^isup>+x. Px x \MX) = 1" + unfolding distributed_distr_eq_density[OF X] using Px + by (subst (asm) emeasure_density) + (auto simp: borel_measurable_ereal_iff intro!: integral_cong cong: positive_integral_cong) + ultimately show False + by (simp add: positive_integral_cong_AE) +qed + +lemma (in information_space) entropy_le: + fixes Px :: "'b \ real" and MX :: "'b measure" + assumes X: "distributed M MX X Px" + and fin: "emeasure MX {x \ space MX. Px x \ 0} \ \" + and int: "integrable MX (\x. - Px x * log b (Px x))" + shows "entropy b MX X \ log b (measure MX {x \ space MX. Px x \ 0})" 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) + note Px = distributed_borel_measurable[OF X] distributed_AE[OF X] + interpret X: prob_space "distr M MX X" + using distributed_measurable[OF X] by (rule prob_space_distr) + + have " - log b (measure MX {x \ space MX. Px x \ 0}) = + - log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX)" + using Px fin + by (subst integral_indicator) (auto simp: measure_def borel_measurable_ereal_iff) + also have "- log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX) = - log b (\ x. 1 / Px x \distr M MX X)" + unfolding distributed_distr_eq_density[OF X] using Px + apply (intro arg_cong[where f="log b"] arg_cong[where f=uminus]) + by (subst integral_density) (auto simp: borel_measurable_ereal_iff intro!: integral_cong) + also have "\ \ (\ x. - log b (1 / Px x) \distr M MX X)" + proof (rule X.jensens_inequality[of "\x. 1 / Px x" "{0<..}" 0 1 "\x. - log b x"]) + show "AE x in distr M MX X. 1 / Px x \ {0<..}" + unfolding distributed_distr_eq_density[OF X] + using Px by (auto simp: AE_density) + have [simp]: "\x. x \ space MX \ ereal (if Px x = 0 then 0 else 1) = indicator {x \ space MX. Px x \ 0} x" + by (auto simp: one_ereal_def) + have "(\\<^isup>+ x. max 0 (ereal (- (if Px x = 0 then 0 else 1))) \MX) = (\\<^isup>+ x. 0 \MX)" + by (intro positive_integral_cong) (auto split: split_max) + then show "integrable (distr M MX X) (\x. 1 / Px x)" + unfolding distributed_distr_eq_density[OF X] using Px + by (auto simp: positive_integral_density integrable_def borel_measurable_ereal_iff fin positive_integral_max_0 + cong: positive_integral_cong) + have "integrable MX (\x. Px x * log b (1 / Px x)) = + integrable MX (\x. - Px x * log b (Px x))" + using Px + by (intro integrable_cong_AE) + (auto simp: borel_measurable_ereal_iff log_divide_eq + intro!: measurable_If) + then show "integrable (distr M MX X) (\x. - log b (1 / Px x))" + unfolding distributed_distr_eq_density[OF X] + using Px int + by (subst integral_density) (auto simp: borel_measurable_ereal_iff) + qed (auto simp: minus_log_convex[OF b_gt_1]) + also have "\ = (\ x. log b (Px x) \distr M MX X)" + unfolding distributed_distr_eq_density[OF X] using Px + by (intro integral_cong_AE) (auto simp: AE_density log_divide_eq) + also have "\ = - entropy b MX X" + unfolding distributed_distr_eq_density[OF X] using Px + by (subst entropy_distr[OF X]) (auto simp: borel_measurable_ereal_iff integral_density) + finally show ?thesis + by simp +qed + +lemma (in information_space) entropy_le_space: + fixes Px :: "'b \ real" and MX :: "'b measure" + assumes X: "distributed M MX X Px" + and fin: "finite_measure MX" + and int: "integrable MX (\x. - Px x * log b (Px x))" + shows "entropy b MX X \ log b (measure MX (space MX))" +proof - + note Px = distributed_borel_measurable[OF X] distributed_AE[OF X] + interpret finite_measure MX by fact + have "entropy b MX X \ log b (measure MX {x \ space MX. Px x \ 0})" + using int X by (intro entropy_le) auto + also have "\ \ log b (measure MX (space MX))" + using Px distributed_imp_emeasure_nonzero[OF X] + by (intro log_le) + (auto intro!: borel_measurable_ereal_iff finite_measure_mono b_gt_1 + less_le[THEN iffD2] measure_nonneg simp: emeasure_eq_measure) + finally show ?thesis . +qed + +lemma (in prob_space) uniform_distributed_params: + assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" + shows "A \ sets MX" "measure MX A \ 0" +proof - + interpret X: prob_space "distr M MX X" + using distributed_measurable[OF X] by (rule prob_space_distr) + + show "measure MX A \ 0" + proof + assume "measure MX A = 0" + with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X] + show False + by (simp add: emeasure_density zero_ereal_def[symmetric]) + qed + with measure_notin_sets[of A MX] show "A \ sets MX" + by blast qed 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)" + assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" (is "distributed _ _ _ ?f") 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+ +proof (subst entropy_distr[OF X]) + have [simp]: "emeasure MX A \ \" + using uniform_distributed_params[OF X] by (auto simp add: measure_def) + have eq: "(\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = + (\ x. (- log b (measure MX A) / measure MX A) * indicator A x \MX)" + using measure_nonneg[of MX A] uniform_distributed_params[OF X] + by (auto intro!: integral_cong split: split_indicator simp: log_divide_eq) + show "- (\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = + log b (measure MX A)" + unfolding eq using uniform_distributed_params[OF X] + by (subst lebesgue_integral_cmult) (auto simp: measure_def) +qed 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 + "simple_distributed M X f \ \(X) = - (\x\X`space M. f x * log b (f x))" + by (subst entropy_distr[OF simple_distributed]) + (auto simp add: lebesgue_integral_count_space_finite) lemma (in information_space) entropy_le_card_not_0: assumes X: "simple_distributed M X f" shows "\(X) \ log b (card (X ` space M \ {x. f x \ 0}))" proof - - 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_distributed M X f` by (auto simp: setsum_cases real_eq_of_nat) + let ?X = "count_space (X`space M)" + have "\(X) \ log b (measure ?X {x \ space ?X. f x \ 0})" + by (rule entropy_le[OF simple_distributed[OF X]]) + (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space) + also have "measure ?X {x \ space ?X. f x \ 0} = card (X ` space M \ {x. f x \ 0})" + by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def Int_def) + finally show ?thesis . qed lemma (in information_space) entropy_le_card: - assumes "simple_distributed M X f" + assumes X: "simple_distributed M X f" shows "\(X) \ log b (real (card (X ` space M)))" -proof cases - 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_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_simple_distributed) -next - 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_distributed_def) +proof - + let ?X = "count_space (X`space M)" + have "\(X) \ log b (measure ?X (space ?X))" + by (rule entropy_le_space[OF simple_distributed[OF X]]) + (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space) + also have "measure ?X (space ?X) = card (X ` space M)" + by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def) finally show ?thesis . qed @@ -757,7 +1007,18 @@ "\(X ; Y | Z) \ conditional_mutual_information b (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: +lemma (in pair_sigma_finite) borel_measurable_positive_integral_fst: + "(\(x, y). f x y) \ borel_measurable (M1 \\<^isub>M M2) \ (\x. \\<^isup>+ y. f x y \M2) \ borel_measurable M1" + using positive_integral_fst_measurable(1)[of "\(x, y). f x y"] by simp + +lemma (in pair_sigma_finite) borel_measurable_positive_integral_snd: + assumes "(\(x, y). f x y) \ borel_measurable (M2 \\<^isub>M M1)" shows "(\x. \\<^isup>+ y. f x y \M1) \ borel_measurable M2" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + from Q.borel_measurable_positive_integral_fst assms show ?thesis by simp +qed + +lemma (in information_space) 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" @@ -766,16 +1027,19 @@ 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))" + shows conditional_mutual_information_generic_eq: "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))" (is "?eq") + and conditional_mutual_information_generic_nonneg: "0 \ conditional_mutual_information b S T P X Y Z" (is "?nonneg") proof - 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 ST: pair_sigma_finite S T .. interpret SPT: pair_sigma_finite "S \\<^isub>M P" T .. interpret STP: pair_sigma_finite S "T \\<^isub>M P" .. + interpret TPS: pair_sigma_finite "T \\<^isub>M P" S .. 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))" @@ -811,27 +1075,27 @@ 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" + have ae1: "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" + moreover have ae2: "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" + moreover have ae3: "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" + moreover have ae4: "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)" + moreover have ae5: "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)" + moreover have ae6: "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))" + moreover have ae7: "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))" + moreover have ae8: "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. + ultimately have ae: "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))) " @@ -846,13 +1110,454 @@ using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps) qed simp qed - with I1 I2 show ?thesis + with I1 I2 show ?eq unfolding conditional_mutual_information_def 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 + + let ?P = "density (S \\<^isub>M T \\<^isub>M P) Pxyz" + interpret P: prob_space ?P + unfolding distributed_distr_eq_density[OF Pxyz, symmetric] + using distributed_measurable[OF Pxyz] by (rule prob_space_distr) + + let ?Q = "density (T \\<^isub>M P) Pyz" + interpret Q: prob_space ?Q + unfolding distributed_distr_eq_density[OF Pyz, symmetric] + using distributed_measurable[OF Pyz] by (rule prob_space_distr) + + let ?f = "\(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" + + from subdensity_real[of snd, OF _ Pyz Pz] + have aeX1: "AE x in T \\<^isub>M P. Pz (snd x) = 0 \ Pyz x = 0" by (auto simp: comp_def) + have aeX2: "AE x in T \\<^isub>M P. 0 \ Pz (snd x)" + using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable) + + have aeX3: "AE y in T \\<^isub>M P. (\\<^isup>+ x. ereal (Pxz (x, snd y)) \S) = ereal (Pz (snd y))" + using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] + apply (intro TP.AE_pair_measure) + apply (auto simp: comp_def measurable_split_conv + intro!: measurable_snd'' borel_measurable_ereal_eq borel_measurable_ereal + SP.borel_measurable_positive_integral_snd measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] + measurable_Pair + dest: distributed_real_AE distributed_real_measurable) + done + + note M = borel_measurable_divide borel_measurable_diff borel_measurable_times borel_measurable_ereal + measurable_compose[OF _ measurable_snd] + measurable_Pair + measurable_compose[OF _ Pxyz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pyz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pz[THEN distributed_real_measurable]] + measurable_compose[OF _ Px[THEN distributed_real_measurable]] + STP.borel_measurable_positive_integral_snd + have "(\\<^isup>+ x. ?f x \?P) \ (\\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \(S \\<^isub>M T \\<^isub>M P))" + apply (subst positive_integral_density) + apply (rule distributed_borel_measurable[OF Pxyz]) + apply (rule distributed_AE[OF Pxyz]) + apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) [] + apply (rule positive_integral_mono_AE) + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg) + done + also have "\ = (\\<^isup>+(y, z). \\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \S \T \\<^isub>M P)" + by (subst STP.positive_integral_snd_measurable[symmetric]) + (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) + also have "\ = (\\<^isup>+x. ereal (Pyz x) * 1 \T \\<^isub>M P)" + apply (rule positive_integral_cong_AE) + using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space + apply eventually_elim + proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure) + fix a b assume "Pz b = 0 \ Pyz (a, b) = 0" "0 \ Pz b" "a \ space T \ b \ space P" + "(\\<^isup>+ x. ereal (Pxz (x, b)) \S) = ereal (Pz b)" "0 \ Pyz (a, b)" + then show "(\\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \S) = ereal (Pyz (a, b))" + apply (subst positive_integral_multc) + apply (auto intro!: borel_measurable_ereal divide_nonneg_nonneg + measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] measurable_Pair + split: prod.split) + done + qed + also have "\ = 1" + using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz] + by (subst positive_integral_density[symmetric]) (auto intro!: M) + finally have le1: "(\\<^isup>+ x. ?f x \?P) \ 1" . + also have "\ < \" by simp + finally have fin: "(\\<^isup>+ x. ?f x \?P) \ \" by simp + + have pos: "(\\<^isup>+ x. ?f x \?P) \ 0" + apply (subst positive_integral_density) + apply (rule distributed_borel_measurable[OF Pxyz]) + apply (rule distributed_AE[OF Pxyz]) + apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) [] + apply (simp add: split_beta') + proof + let ?g = "\x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))" + assume "(\\<^isup>+ x. ?g x \(S \\<^isub>M T \\<^isub>M P)) = 0" + then have "AE x in S \\<^isub>M T \\<^isub>M P. ?g x \ 0" + by (intro positive_integral_0_iff_AE[THEN iffD1]) (auto intro!: M borel_measurable_ereal measurable_If) + then have "AE x in S \\<^isub>M T \\<^isub>M P. Pxyz x = 0" + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff) + then have "(\\<^isup>+ x. ereal (Pxyz x) \S \\<^isub>M T \\<^isub>M P) = 0" + by (subst positive_integral_cong_AE[of _ "\x. 0"]) auto + with P.emeasure_space_1 show False + by (subst (asm) emeasure_density) (auto intro!: M cong: positive_integral_cong) + qed + + have neg: "(\\<^isup>+ x. - ?f x \?P) = 0" + apply (rule positive_integral_0_iff_AE[THEN iffD2]) + apply (auto intro!: M simp: split_beta') [] + apply (subst AE_density) + apply (auto intro!: M simp: split_beta') [] + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) + done + + have I3: "integrable (S \\<^isub>M T \\<^isub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]]) + using ae + apply (auto intro!: M simp: split_beta') + done + + have "- log b 1 \ - log b (integral\<^isup>L ?P ?f)" + proof (intro le_imp_neg_le log_le[OF b_gt_1]) + show "0 < integral\<^isup>L ?P ?f" + using neg pos fin positive_integral_positive[of ?P ?f] + by (cases "(\\<^isup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def less_le split_beta') + show "integral\<^isup>L ?P ?f \ 1" + using neg le1 fin positive_integral_positive[of ?P ?f] + by (cases "(\\<^isup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def) + qed + also have "- log b (integral\<^isup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" + proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) + show "AE x in ?P. ?f x \ {0<..}" + unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + by eventually_elim (auto simp: divide_pos_pos mult_pos_pos) + show "integrable ?P ?f" + unfolding integrable_def + using fin neg by (auto intro!: M simp: split_beta') + show "integrable ?P (\x. - log b (?f x))" + apply (subst integral_density) + apply (auto intro!: M) [] + apply (auto intro!: M distributed_real_AE[OF Pxyz]) [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + apply eventually_elim + apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps) + done + qed (auto simp: b_gt_1 minus_log_convex) + also have "\ = conditional_mutual_information b S T P X Y Z" + unfolding `?eq` + apply (subst integral_density) + apply (auto intro!: M) [] + apply (auto intro!: M distributed_real_AE[OF Pxyz]) [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (intro integral_cong_AE) + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + apply eventually_elim + apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps) + done + finally show ?nonneg + by simp +qed + +lemma (in information_space) + fixes Px :: "_ \ real" + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" + assumes Fx: "finite_entropy S X Px" + assumes Fz: "finite_entropy P Z Pz" + assumes Fyz: "finite_entropy (T \\<^isub>M P) (\x. (Y x, Z x)) Pyz" + assumes Fxz: "finite_entropy (S \\<^isub>M P) (\x. (X x, Z x)) Pxz" + assumes Fxyz: "finite_entropy (S \\<^isub>M T \\<^isub>M P) (\x. (X x, Y x, Z x)) Pxyz" + shows conditional_mutual_information_generic_eq': "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))" (is "?eq") + and conditional_mutual_information_generic_nonneg': "0 \ conditional_mutual_information b S T P X Y Z" (is "?nonneg") +proof - + note Px = Fx[THEN finite_entropy_distributed] + note Pz = Fz[THEN finite_entropy_distributed] + note Pyz = Fyz[THEN finite_entropy_distributed] + note Pxz = Fxz[THEN finite_entropy_distributed] + note Pxyz = Fxyz[THEN finite_entropy_distributed] + + 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 ST: pair_sigma_finite S T .. + interpret SPT: pair_sigma_finite "S \\<^isub>M P" T .. + interpret STP: pair_sigma_finite S "T \\<^isub>M P" .. + interpret TPS: pair_sigma_finite "T \\<^isub>M P" S .. + 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) + + 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 ae1: "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 ae2: "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 ae3: "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 ae4: "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 ae5: "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 ae6: "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 ae7: "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 ae8: "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 ae9 = Pxyz[THEN distributed_real_AE] + ultimately have ae: "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 "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 + + have "integrable (S \\<^isub>M T \\<^isub>M P) + (\x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))" + using finite_entropy_integrable[OF Fxyz] + using finite_entropy_integrable_transform[OF Fx Pxyz, of fst] + using finite_entropy_integrable_transform[OF Fyz Pxyz, of snd] + by simp + moreover have "(\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \ borel_measurable (S \\<^isub>M T \\<^isub>M P)" + using Pxyz Px Pyz + by (auto intro!: borel_measurable_times measurable_fst'' measurable_snd'' dest!: distributed_real_measurable simp: split_beta') + ultimately have 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))))" + apply (rule integrable_cong_AE_imp) + using ae1 ae4 ae5 ae6 ae9 + by eventually_elim + (auto simp: log_divide_eq log_mult_eq mult_nonneg_nonneg field_simps zero_less_mult_iff) + + have "integrable (S \\<^isub>M T \\<^isub>M P) + (\x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))" + using finite_entropy_integrable_transform[OF Fxz Pxyz, of "\x. (fst x, snd (snd x))"] + using finite_entropy_integrable_transform[OF Fx Pxyz, of fst] + using finite_entropy_integrable_transform[OF Fz Pxyz, of "snd \ snd"] + by (simp add: measurable_Pair measurable_snd'' comp_def) + moreover have "(\(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \ borel_measurable (S \\<^isub>M T \\<^isub>M P)" + using Pxyz Px Pz + by (auto intro!: measurable_compose[OF _ distributed_real_measurable[OF Pxz]] + measurable_Pair borel_measurable_times measurable_fst'' measurable_snd'' + dest!: distributed_real_measurable simp: split_beta') + ultimately have 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)))" + apply (rule integrable_cong_AE_imp) + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 ae9 + by eventually_elim + (auto simp: log_divide_eq log_mult_eq mult_nonneg_nonneg field_simps zero_less_mult_iff) + + from ae I1 I2 show ?eq + unfolding conditional_mutual_information_def + 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 + + let ?P = "density (S \\<^isub>M T \\<^isub>M P) Pxyz" + interpret P: prob_space ?P + unfolding distributed_distr_eq_density[OF Pxyz, symmetric] + using distributed_measurable[OF Pxyz] by (rule prob_space_distr) + + let ?Q = "density (T \\<^isub>M P) Pyz" + interpret Q: prob_space ?Q + unfolding distributed_distr_eq_density[OF Pyz, symmetric] + using distributed_measurable[OF Pyz] by (rule prob_space_distr) + + let ?f = "\(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" + + from subdensity_real[of snd, OF _ Pyz Pz] + have aeX1: "AE x in T \\<^isub>M P. Pz (snd x) = 0 \ Pyz x = 0" by (auto simp: comp_def) + have aeX2: "AE x in T \\<^isub>M P. 0 \ Pz (snd x)" + using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable) + + have aeX3: "AE y in T \\<^isub>M P. (\\<^isup>+ x. ereal (Pxz (x, snd y)) \S) = ereal (Pz (snd y))" + using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] + apply (intro TP.AE_pair_measure) + apply (auto simp: comp_def measurable_split_conv + intro!: measurable_snd'' borel_measurable_ereal_eq borel_measurable_ereal + SP.borel_measurable_positive_integral_snd measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] + measurable_Pair + dest: distributed_real_AE distributed_real_measurable) + done + + note M = borel_measurable_divide borel_measurable_diff borel_measurable_times borel_measurable_ereal + measurable_compose[OF _ measurable_snd] + measurable_Pair + measurable_compose[OF _ Pxyz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pyz[THEN distributed_real_measurable]] + measurable_compose[OF _ Pz[THEN distributed_real_measurable]] + measurable_compose[OF _ Px[THEN distributed_real_measurable]] + STP.borel_measurable_positive_integral_snd + have "(\\<^isup>+ x. ?f x \?P) \ (\\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \(S \\<^isub>M T \\<^isub>M P))" + apply (subst positive_integral_density) + apply (rule distributed_borel_measurable[OF Pxyz]) + apply (rule distributed_AE[OF Pxyz]) + apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) [] + apply (rule positive_integral_mono_AE) + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg) + done + also have "\ = (\\<^isup>+(y, z). \\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \S \T \\<^isub>M P)" + by (subst STP.positive_integral_snd_measurable[symmetric]) + (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) + also have "\ = (\\<^isup>+x. ereal (Pyz x) * 1 \T \\<^isub>M P)" + apply (rule positive_integral_cong_AE) + using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space + apply eventually_elim + proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure) + fix a b assume "Pz b = 0 \ Pyz (a, b) = 0" "0 \ Pz b" "a \ space T \ b \ space P" + "(\\<^isup>+ x. ereal (Pxz (x, b)) \S) = ereal (Pz b)" "0 \ Pyz (a, b)" + then show "(\\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \S) = ereal (Pyz (a, b))" + apply (subst positive_integral_multc) + apply (auto intro!: borel_measurable_ereal divide_nonneg_nonneg + measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] measurable_Pair + split: prod.split) + done + qed + also have "\ = 1" + using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz] + by (subst positive_integral_density[symmetric]) (auto intro!: M) + finally have le1: "(\\<^isup>+ x. ?f x \?P) \ 1" . + also have "\ < \" by simp + finally have fin: "(\\<^isup>+ x. ?f x \?P) \ \" by simp + + have pos: "(\\<^isup>+ x. ?f x \?P) \ 0" + apply (subst positive_integral_density) + apply (rule distributed_borel_measurable[OF Pxyz]) + apply (rule distributed_AE[OF Pxyz]) + apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) [] + apply (simp add: split_beta') + proof + let ?g = "\x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))" + assume "(\\<^isup>+ x. ?g x \(S \\<^isub>M T \\<^isub>M P)) = 0" + then have "AE x in S \\<^isub>M T \\<^isub>M P. ?g x \ 0" + by (intro positive_integral_0_iff_AE[THEN iffD1]) (auto intro!: M borel_measurable_ereal measurable_If) + then have "AE x in S \\<^isub>M T \\<^isub>M P. Pxyz x = 0" + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff) + then have "(\\<^isup>+ x. ereal (Pxyz x) \S \\<^isub>M T \\<^isub>M P) = 0" + by (subst positive_integral_cong_AE[of _ "\x. 0"]) auto + with P.emeasure_space_1 show False + by (subst (asm) emeasure_density) (auto intro!: M cong: positive_integral_cong) + qed + + have neg: "(\\<^isup>+ x. - ?f x \?P) = 0" + apply (rule positive_integral_0_iff_AE[THEN iffD2]) + apply (auto intro!: M simp: split_beta') [] + apply (subst AE_density) + apply (auto intro!: M simp: split_beta') [] + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) + done + + have I3: "integrable (S \\<^isub>M T \\<^isub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]]) + using ae + apply (auto intro!: M simp: split_beta') + done + + have "- log b 1 \ - log b (integral\<^isup>L ?P ?f)" + proof (intro le_imp_neg_le log_le[OF b_gt_1]) + show "0 < integral\<^isup>L ?P ?f" + using neg pos fin positive_integral_positive[of ?P ?f] + by (cases "(\\<^isup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def less_le split_beta') + show "integral\<^isup>L ?P ?f \ 1" + using neg le1 fin positive_integral_positive[of ?P ?f] + by (cases "(\\<^isup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def) + qed + also have "- log b (integral\<^isup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" + proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) + show "AE x in ?P. ?f x \ {0<..}" + unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + by eventually_elim (auto simp: divide_pos_pos mult_pos_pos) + show "integrable ?P ?f" + unfolding integrable_def + using fin neg by (auto intro!: M simp: split_beta') + show "integrable ?P (\x. - log b (?f x))" + apply (subst integral_density) + apply (auto intro!: M) [] + apply (auto intro!: M distributed_real_AE[OF Pxyz]) [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + apply eventually_elim + apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps) + done + qed (auto simp: b_gt_1 minus_log_convex) + also have "\ = conditional_mutual_information b S T P X Y Z" + unfolding `?eq` + apply (subst integral_density) + apply (auto intro!: M) [] + apply (auto intro!: M distributed_real_AE[OF Pxyz]) [] + apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') [] + apply (intro integral_cong_AE) + using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] + apply eventually_elim + apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps) + done + finally show ?nonneg + by simp qed lemma (in information_space) conditional_mutual_information_eq: @@ -909,92 +1614,28 @@ 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 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) - - { 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 - - { 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 not_empty by simp - show "1 < b" using b_gt_1 . - - show "finite ?M" using X Y Z by (auto simp: simple_functionD) - - 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. 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) - apply (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] marginal1 marginal2 - intro!: setsum_cong) - 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 + have [simp]: "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)" + by (simp add: pair_measure_count_space X Y Z simple_functionD) + note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)] + note sd = simple_distributedI[OF _ refl] + note sp = simple_function_Pair + show ?thesis + apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]]) + apply (rule simple_distributed[OF sd[OF X]]) + apply (rule simple_distributed[OF sd[OF Z]]) + apply (rule simple_distributed_joint[OF sd[OF sp[OF Y Z]]]) + apply (rule simple_distributed_joint[OF sd[OF sp[OF X Z]]]) + apply (rule simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]]) + apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD) + done qed subsection {* Conditional Entropy *} definition (in prob_space) - "conditional_entropy b S T X Y = entropy b (S \\<^isub>M T) (\x. (X x, Y x)) - entropy b T Y" + "conditional_entropy b S T X Y = - (\(x, y). log b (real (RN_deriv (S \\<^isub>M T) (distr M (S \\<^isub>M T) (\x. (X x, Y x))) (x, y)) / + real (RN_deriv T (distr M T Y) y)) \distr M (S \\<^isub>M T) (\x. (X x, Y x)))" abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where @@ -1003,66 +1644,112 @@ lemma (in information_space) conditional_entropy_generic_eq: 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 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) + have "AE x in density (S \\<^isub>M T) (\x. ereal (Pxy x)). Pxy x = real (RN_deriv (S \\<^isub>M T) (distr M (S \\<^isub>M T) (\x. (X x, Y x))) x)" + unfolding AE_density[OF distributed_borel_measurable, OF Pxy] + unfolding distributed_distr_eq_density[OF Pxy] + using distributed_RN_deriv[OF Pxy] + by auto + moreover + have "AE x in density (S \\<^isub>M T) (\x. ereal (Pxy x)). Py (snd x) = real (RN_deriv T (distr M T Y) (snd x))" + unfolding AE_density[OF distributed_borel_measurable, OF Pxy] + unfolding distributed_distr_eq_density[OF Py] + apply (rule ST.AE_pair_measure) + apply (auto intro!: sets_Collect borel_measurable_eq measurable_compose[OF _ distributed_real_measurable[OF Py]] + distributed_real_measurable[OF Pxy] distributed_real_AE[OF Py] + borel_measurable_real_of_ereal measurable_compose[OF _ borel_measurable_RN_deriv_density]) + using distributed_RN_deriv[OF Py] + apply auto + done + ultimately + have "conditional_entropy b S T X Y = - (\x. Pxy x * log b (Pxy x / Py (snd x)) \(S \\<^isub>M T))" + unfolding conditional_entropy_def neg_equal_iff_equal + apply (subst integral_density(1)[symmetric]) + apply (auto simp: distributed_real_measurable[OF Pxy] distributed_real_AE[OF Pxy] + measurable_compose[OF _ distributed_real_measurable[OF Py]] + distributed_distr_eq_density[OF Pxy] + intro!: integral_cong_AE) + done + then show ?thesis by (simp add: split_beta') +qed - 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) +lemma (in information_space) conditional_entropy_eq_entropy: + fixes Px :: "'b \ real" and Py :: "'c \ real" + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + 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 = entropy b (S \\<^isub>M T) (\x. (X x, Y x)) - entropy b T Y" +proof - + interpret S: sigma_finite_measure S by fact + interpret T: sigma_finite_measure T by fact + interpret ST: pair_sigma_finite S T .. have "entropy b T Y = - (\y. Py y * log b (Py y) \T)" - by (rule entropy_distr[OF T Py]) + by (rule entropy_distr[OF 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" + + have ae2: "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)" + moreover have ae4: "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)))" + moreover note ae5 = Pxy[THEN distributed_real_AE] + ultimately have "AE x in S \\<^isub>M T. 0 \ Pxy x \ 0 \ Py (snd x) \ + (Pxy x = 0 \ (Pxy x \ 0 \ 0 < Pxy x \ 0 < Py (snd x)))" by eventually_elim auto - - from pos have "AE x in S \\<^isub>M T. + then have ae: "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 - 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) + have "conditional_entropy b S T X Y = + - (\x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \(S \\<^isub>M T))" + unfolding conditional_entropy_generic_eq[OF S T Py Pxy] neg_equal_iff_equal + apply (intro integral_cong_AE) + using ae + apply eventually_elim + apply auto done + also have "\ = - (\x. Pxy x * log b (Pxy x) \(S \\<^isub>M T)) - - (\x. Pxy x * log b (Py (snd x)) \(S \\<^isub>M T))" + by (simp add: integral_diff[OF I1 I2]) + finally show ?thesis + unfolding conditional_entropy_generic_eq[OF S T Py Pxy] entropy_distr[OF Pxy] e_eq + by (simp add: split_beta') +qed + +lemma (in information_space) conditional_entropy_eq_entropy_simple: + assumes X: "simple_function M X" and Y: "simple_function M Y" + shows "\(X | Y) = entropy b (count_space (X`space M) \\<^isub>M count_space (Y`space M)) (\x. (X x, Y x)) - \(Y)" +proof - + 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_functionD pair_measure_count_space) + show ?thesis + by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite + simple_functionD X Y simple_distributed simple_distributedI[OF _ refl] + simple_distributed_joint simple_function_Pair integrable_count_space)+ + (auto simp: `?P = ?C` intro!: integrable_count_space simple_functionD X Y) qed lemma (in information_space) conditional_entropy_eq: - assumes Y: "simple_distributed M Y Py" and X: "simple_function M X" + assumes 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) / 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) + simple_distributed[OF Y] simple_distributed_joint[OF XY]]) + have "finite ((\x. (X x, Y x))`space M)" + using XY unfolding simple_distributed_def by auto + from finite_imageI[OF this, of fst] + have [simp]: "finite (X`space M)" + by (simp add: image_compose[symmetric] comp_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) @@ -1071,15 +1758,11 @@ 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) + using Y by (simp add: simple_distributed_finite pair_measure_count_space) 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) = + from 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 @@ -1111,11 +1794,11 @@ 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 (subst conditional_entropy_eq[OF Py 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) + apply (auto simp add: not_le[symmetric] AE_count_space) done qed @@ -1138,14 +1821,14 @@ proof - 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 entropy_distr[OF 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 entropy_distr[OF Py]) apply (subst distributed_transform_integral[OF Pxy Py, where T=snd]) apply (auto intro!: integral_cong) done @@ -1156,7 +1839,7 @@ 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) + by (subst entropy_distr[OF 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) @@ -1197,6 +1880,20 @@ finally show ?thesis .. qed +lemma (in information_space) mutual_information_eq_entropy_conditional_entropy': + 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 - conditional_entropy b S T X Y" + using + mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy] + conditional_entropy_eq_entropy[OF S T Py Pxy Ixy Iy] + by simp + 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)" @@ -1217,7 +1914,7 @@ 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 + unfolding conditional_entropy_eq_entropy_simple[OF sf_X sf_Y] by simp qed lemma (in information_space) mutual_information_nonneg_simple: @@ -1252,6 +1949,25 @@ finally show ?thesis by auto qed +lemma (in information_space) + 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: "finite_entropy S X Px" and Py: "finite_entropy T Y Py" + assumes Pxy: "finite_entropy (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "conditional_entropy b S T X Y \ entropy b S X" +proof - + + have "0 \ mutual_information b S T X Y" + by (rule mutual_information_nonneg') fact+ + also have "\ = entropy b S X - conditional_entropy b S T X Y" + apply (rule mutual_information_eq_entropy_conditional_entropy') + using assms + by (auto intro!: finite_entropy_integrable finite_entropy_distributed + finite_entropy_integrable_transform[OF Px] + finite_entropy_integrable_transform[OF Py]) + 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)" @@ -1269,12 +1985,12 @@ 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]]) + 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 + unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp qed lemma (in information_space) entropy_partition: diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Wed Oct 10 13:30:50 2012 +0200 @@ -355,6 +355,98 @@ "\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 simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. P u \ P (\x. c * u x)" + assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" + shows "P u" +proof (rule cong) + from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" + proof eventually_elim + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + qed +next + from u have "finite (u ` space M)" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + qed (auto intro!: add mult set simple_functionD u) +next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done +qed fact + +lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" and nn: "\x. 0 \ u x" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. simple_function M u \ (\x. 0 \ u x) \ P u \ simple_function M v \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" + shows "P u" +proof - + show ?thesis + proof (rule cong) + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done + next + from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + qed (auto intro!: add mult set simple_functionD u setsum_nonneg + simple_function_setsum) + qed fact +qed + +lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" + assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ P (SUP i. U i)" + shows "P u" + using u +proof (induct rule: borel_measurable_implies_simple_function_sequence') + fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and + sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" + have u_eq: "u = (SUP i. U i)" + using nn u sup by (auto simp: max_def) + + from U have "\i. U i \ borel_measurable M" + by (simp add: borel_measurable_simple_function) + + show "P u" + unfolding u_eq + proof (rule seq) + fix i show "P (U i)" + using `simple_function M (U i)` nn + by (induct rule: simple_function_induct_nn) + (auto intro: set mult add cong dest!: borel_measurable_simple_function) + qed fact+ +qed + 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") @@ -595,7 +687,7 @@ lemma simple_integral_indicator: assumes "A \ sets M" - assumes "simple_function M f" + assumes f: "simple_function M f" shows "(\\<^isup>Sx. f x * indicator A x \M) = (\x \ f ` space M. x * (emeasure M) (f -` {x} \ space M \ A))" proof cases @@ -670,32 +762,6 @@ shows "integral\<^isup>S M f = (\\<^isup>Sx. f x * indicator S x \M)" using assms by (intro simple_integral_cong_AE) auto -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 (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 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 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 simple_integral_cmult_indicator: assumes A: "A \ sets M" shows "(\\<^isup>Sx. c * indicator A x \M) = c * (emeasure M) A" @@ -1088,7 +1154,7 @@ qed lemma positive_integral_cmult: - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ c" + assumes f: "f \ borel_measurable M" "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` @@ -1101,14 +1167,14 @@ qed lemma positive_integral_multc: - assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ c" + assumes "f \ borel_measurable M" "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 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) + (auto simp: simple_integral_indicator) lemma positive_integral_cmult_indicator: "0 \ c \ A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * (emeasure M) A" @@ -1151,7 +1217,7 @@ qed simp 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 \ \" + assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ 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 - @@ -1342,23 +1408,21 @@ by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) lemma positive_integral_subalgebra: - assumes f: "f \ borel_measurable N" "AE x in N. 0 \ f x" + assumes f: "f \ borel_measurable N" "\x. 0 \ f x" 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 - - 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 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 * (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 - unfolding simple_integral_def `space N = space M` by simp - finally show ?thesis . + have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" + using N by (auto simp: measurable_def) + have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" + using N by (auto simp add: eventually_ae_filter null_sets_def) + have [simp]: "\A. A \ sets N \ A \ sets M" + using N by auto + from f show ?thesis + apply induct + apply (simp_all add: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N) + apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric]) + done qed section "Lebesgue Integral" @@ -1423,6 +1487,25 @@ "(\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 integral_mono_AE: + assumes fg: "integrable M f" "integrable M g" + 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 in M. ereal (f x) \ ereal (g x)" + using mono by auto + 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 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 positive_integral_eq_integral: assumes f: "integrable M f" assumes nonneg: "AE x in M. 0 \ f x" @@ -1571,20 +1654,16 @@ 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))))" + { have "real (ereal c * integral\<^isup>P M (\x. max 0 (ereal (f x)))) = + 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))))" + { have "real (ereal c * integral\<^isup>P M (\x. max 0 (ereal (- f x)))) = + 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) @@ -1615,25 +1694,6 @@ shows "(\ x. f x * c \M) = integral\<^isup>L M f * c" unfolding mult_commute[of _ c] integral_cmult[OF assms] .. -lemma integral_mono_AE: - assumes fg: "integrable M f" "integrable M g" - 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 in M. ereal (f x) \ ereal (g x)" - using mono by auto - 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 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 integral_diff[simp, intro]: assumes f: "integrable M f" and g: "integrable M g" shows "integrable M (\t. f t - g t)" @@ -1685,8 +1745,33 @@ thus "?int S" and "?I S" by auto qed +lemma integrable_bound: + assumes "integrable M f" and f: "AE x in M. \g x\ \ f x" + assumes borel: "g \ borel_measurable M" + shows "integrable M g" +proof - + have "(\\<^isup>+ x. ereal (g x) \M) \ (\\<^isup>+ x. ereal \g x\ \M)" + by (auto intro!: positive_integral_mono) + also have "\ \ (\\<^isup>+ x. ereal (f x) \M)" + using f by (auto intro!: positive_integral_mono_AE) + also have "\ < \" + using `integrable M f` unfolding integrable_def by auto + finally have pos: "(\\<^isup>+ x. ereal (g x) \M) < \" . + + have "(\\<^isup>+ x. ereal (- g x) \M) \ (\\<^isup>+ x. ereal (\g x\) \M)" + by (auto intro!: positive_integral_mono) + also have "\ \ (\\<^isup>+ x. ereal (f x) \M)" + using f by (auto intro!: positive_integral_mono_AE) + also have "\ < \" + using `integrable M f` unfolding integrable_def by auto + finally have neg: "(\\<^isup>+ x. ereal (- g x) \M) < \" . + + from neg pos borel show ?thesis + unfolding integrable_def by auto +qed + lemma integrable_abs: - assumes "integrable M f" + assumes f: "integrable M f" shows "integrable M (\ x. \f x\)" proof - from assms have *: "(\\<^isup>+x. ereal (- \f x\)\M) = 0" @@ -1711,33 +1796,6 @@ by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0) qed -lemma integrable_bound: - assumes "integrable M f" - and f: "\x. x \ space M \ 0 \ f x" - "\x. x \ space M \ \g x\ \ f x" - assumes borel: "g \ borel_measurable M" - shows "integrable M g" -proof - - have "(\\<^isup>+ x. ereal (g x) \M) \ (\\<^isup>+ x. ereal \g x\ \M)" - by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. ereal (f x) \M)" - using f by (auto intro!: positive_integral_mono) - also have "\ < \" - using `integrable M f` unfolding integrable_def by auto - finally have pos: "(\\<^isup>+ x. ereal (g x) \M) < \" . - - have "(\\<^isup>+ x. ereal (- g x) \M) \ (\\<^isup>+ x. ereal (\g x\) \M)" - by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. ereal (f x) \M)" - using f by (auto intro!: positive_integral_mono) - also have "\ < \" - using `integrable M f` unfolding integrable_def by auto - finally have neg: "(\\<^isup>+ x. ereal (- g x) \M) < \" . - - from neg pos borel show ?thesis - unfolding integrable_def by auto -qed - lemma lebesgue_integral_nonneg: assumes ae: "(AE x in M. 0 \ f x)" shows "0 \ integral\<^isup>L M f" proof - @@ -1760,11 +1818,7 @@ using int by (simp add: integrable_abs) show "(\x. max (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto -next - fix x assume "x \ space M" - show "0 \ \f x\ + \g x\" "\max (f x) (g x)\ \ \f x\ + \g x\" - by auto -qed +qed auto lemma integrable_min: assumes int: "integrable M f" "integrable M g" @@ -1774,11 +1828,7 @@ using int by (simp add: integrable_abs) show "(\x. min (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto -next - fix x assume "x \ space M" - show "0 \ \f x\ + \g x\" "\min (f x) (g x)\ \ \f x\ + \g x\" - by auto -qed +qed auto lemma integral_triangle_inequality: assumes "integrable M f" @@ -1802,74 +1852,71 @@ qed 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" - and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" + and pos: "\i. AE x in M. 0 \ f i x" + and lim: "AE x in M. (\i. f i x) ----> u x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^isup>L M u = x" proof - - { fix x have "0 \ u x" - using mono pos[of 0 x] incseq_le[OF _ lim, of x 0] - by (simp add: mono_def incseq_def) } - note pos_u = this - - have SUP_F: "\x. (SUP n. ereal (f n x)) = ereal (u x)" - unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim) - - have borel_f: "\i. (\x. ereal (f i x)) \ borel_measurable M" - using i unfolding integrable_def by auto - hence "(\x. SUP i. ereal (f i x)) \ borel_measurable M" - by auto - hence borel_u: "u \ borel_measurable M" - by (auto simp: borel_measurable_ereal_iff SUP_F) - - hence [simp]: "\i. (\\<^isup>+x. ereal (- f i x) \M) = 0" "(\\<^isup>+x. ereal (- u x) \M) = 0" - 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[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) - hence "0 \ x" - using LIMSEQ_le_const[OF ilim, of 0] by auto - - from mono pos i have pI: "(\\<^isup>+ x. ereal (u x) \M) = (SUP n. (\\<^isup>+ x. ereal (f n x) \M))" - by (auto intro!: positive_integral_monotone_convergence_SUP - simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric]) - also have "\ = ereal x" unfolding integral_eq - proof (rule SUP_eq_LIMSEQ[THEN iffD2]) - show "mono (\n. integral\<^isup>L M (f n))" - using mono i by (auto simp: mono_def intro!: integral_mono) - show "(\n. integral\<^isup>L M (f n)) ----> x" using ilim . + have "(\\<^isup>+ x. ereal (u x) \M) = (SUP n. (\\<^isup>+ x. ereal (f n x) \M))" + proof (subst positive_integral_monotone_convergence_SUP_AE[symmetric]) + fix i + from mono pos show "AE x in M. ereal (f i x) \ ereal (f (Suc i) x) \ 0 \ ereal (f i x)" + by eventually_elim (auto simp: mono_def) + show "(\x. ereal (f i x)) \ borel_measurable M" + using i by (auto simp: integrable_def) + next + show "(\\<^isup>+ x. ereal (u x) \M) = \\<^isup>+ x. (SUP i. ereal (f i x)) \M" + apply (rule positive_integral_cong_AE) + using lim mono + by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2]) qed - finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \ x` - unfolding integrable_def lebesgue_integral_def by auto + also have "\ = ereal x" + using mono i unfolding positive_integral_eq_integral[OF i pos] + by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim) + finally have "(\\<^isup>+ x. ereal (u x) \M) = ereal x" . + moreover have "(\\<^isup>+ x. ereal (- u x) \M) = 0" + proof (subst positive_integral_0_iff_AE) + show "(\x. ereal (- u x)) \ borel_measurable M" + using u by auto + from mono pos[of 0] lim show "AE x in M. ereal (- u x) \ 0" + proof eventually_elim + fix x assume "mono (\n. f n x)" "0 \ f 0 x" "(\i. f i x) ----> u x" + then show "ereal (- u x) \ 0" + using incseq_le[of "\n. f n x" "u x" 0] by (simp add: mono_def incseq_def) + qed + qed + ultimately show "integrable M u" "integral\<^isup>L M u = x" + by (auto simp: integrable_def lebesgue_integral_def u) qed lemma integral_monotone_convergence: - assumes f: "\i. integrable M (f i)" and "mono f" - and lim: "\x. (\i. f i x) ----> u x" + assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" + and lim: "AE x in M. (\i. f i x) ----> u x" and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^isup>L M u = x" proof - have 1: "\i. integrable M (\x. f i x - f 0 x)" - using f by (auto intro!: integral_diff) - have 2: "\x. mono (\n. f n x - f 0 x)" using `mono f` - unfolding mono_def le_fun_def by auto - have 3: "\x n. 0 \ f n x - f 0 x" using `mono f` - unfolding mono_def le_fun_def by (auto simp: field_simps) - have 4: "\x. (\i. f i x - f 0 x) ----> u x - f 0 x" + using f by auto + have 2: "AE x in M. mono (\n. f n x - f 0 x)" + using mono by (auto simp: mono_def le_fun_def) + have 3: "\n. AE x in M. 0 \ f n x - f 0 x" + using mono by (auto simp: field_simps mono_def le_fun_def) + have 4: "AE x in M. (\i. f i x - f 0 x) ----> u x - f 0 x" using lim by (auto intro!: tendsto_diff) have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^isup>L M (f 0)" - using f ilim by (auto intro!: tendsto_diff simp: integral_diff) - note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5] + using f ilim by (auto intro!: tendsto_diff) + have 6: "(\x. u x - f 0 x) \ borel_measurable M" + using f[of 0] u by auto + note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5 6] have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integral_add(1)) with diff(2) f show "integrable M u" "integral\<^isup>L M u = x" - by (auto simp: integral_diff) + by auto qed lemma integral_0_iff: @@ -1993,58 +2040,47 @@ using gt by (intro integral_less_AE[OF int, where A="space M"]) auto lemma integral_dominated_convergence: - assumes u: "\i. integrable M (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" + assumes u: "\i. integrable M (u i)" and bound: "\j. AE x in M. \u j x\ \ w x" and w: "integrable M w" - and u': "\x. x \ space M \ (\i. u i x) ----> u' x" + and u': "AE x in M. (\i. u i x) ----> u' x" + and borel: "u' \ borel_measurable M" shows "integrable M u'" and "(\i. (\x. \u i x - u' x\ \M)) ----> 0" (is "?lim_diff") and "(\i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim) proof - - { fix x j assume x: "x \ space M" - from u'[OF x] have "(\i. \u i x\) ----> \u' x\" by (rule tendsto_rabs) - from LIMSEQ_le_const2[OF this] - have "\u' x\ \ w x" using bound[OF x] by auto } - note u'_bound = this + have all_bound: "AE x in M. \j. \u j x\ \ w x" + using bound by (auto simp: AE_all_countable) + with u' have u'_bound: "AE x in M. \u' x\ \ w x" + by eventually_elim (auto intro: LIMSEQ_le_const2 tendsto_rabs) - from u[unfolded integrable_def] - have u'_borel: "u' \ borel_measurable M" - 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 - also have "\ \ w x" using bound[OF x] by auto - finally have "0 \ w x" . } - note w_pos = this + from bound[of 0] have w_pos: "AE x in M. 0 \ w x" + by eventually_elim auto show "integrable M u'" - proof (rule integrable_bound) - show "integrable M w" by fact - show "u' \ borel_measurable M" by fact - next - fix x assume x: "x \ space M" then show "0 \ w x" by fact - show "\u' x\ \ w x" using u'_bound[OF x] . - qed + by (rule integrable_bound) fact+ let ?diff = "\n x. 2 * w x - \u n x - u' x\" have diff: "\n. integrable M (\x. \u n x - u' x\)" - using w u `integrable M u'` - by (auto intro!: integral_add integral_diff integral_cmult integrable_abs) + using w u `integrable M u'` by (auto intro!: integrable_abs) - { fix j x assume x: "x \ space M" - have "\u j x - u' x\ \ \u j x\ + \u' x\" by auto + from u'_bound all_bound + have diff_less_2w: "AE x in M. \j. \u j x - u' x\ \ 2 * w x" + proof (eventually_elim, intro allI) + fix x j assume *: "\u' x\ \ w x" "\j. \u j x\ \ w x" + then have "\u j x - u' x\ \ \u j x\ + \u' x\" by auto also have "\ \ w x + w x" - by (rule add_mono[OF bound[OF x] u'_bound[OF x]]) - finally have "\u j x - u' x\ \ 2 * w x" by simp } - note diff_less_2w = this + using * by (intro add_mono) auto + finally show "\u j x - u' x\ \ 2 * w x" by simp + qed have PI_diff: "\n. (\\<^isup>+ x. ereal (?diff n x) \M) = (\\<^isup>+ x. ereal (2 * w x) \M) - (\\<^isup>+ x. ereal \u n x - u' x\ \M)" using diff w diff_less_2w w_pos by (subst positive_integral_diff[symmetric]) - (auto simp: integrable_def intro!: positive_integral_cong) + (auto simp: integrable_def intro!: positive_integral_cong_AE) have "integrable M (\x. 2 * w x)" - using w by (auto intro: integral_cmult) + using w by auto hence I2w_fin: "(\\<^isup>+ x. ereal (2 * w x) \M) \ \" and borel_2w: "(\x. ereal (2 * w x)) \ borel_measurable M" unfolding integrable_def by auto @@ -2054,8 +2090,8 @@ assume eq_0: "(\\<^isup>+ x. max 0 (ereal (2 * w x)) \M) = 0" (is "?wx = 0") { fix n have "?f n \ ?wx" (is "integral\<^isup>P M ?f' \ _") - using diff_less_2w[of _ n] unfolding positive_integral_max_0 - by (intro positive_integral_mono) auto + using diff_less_2w unfolding positive_integral_max_0 + by (intro positive_integral_mono_AE) auto then have "?f n = 0" using positive_integral_positive[of M ?f'] eq_0 by auto } then show ?thesis by (simp add: Limsup_const) @@ -2066,19 +2102,20 @@ by (intro limsup_mono positive_integral_positive) finally have pos: "0 \ limsup (\n. \\<^isup>+ x. ereal \u n x - u' x\ \M)" . have "?wx = (\\<^isup>+ x. liminf (\n. max 0 (ereal (?diff n x))) \M)" - proof (rule positive_integral_cong) - fix x assume x: "x \ space M" + using u' + proof (intro positive_integral_cong_AE, eventually_elim) + fix x assume u': "(\i. u i x) ----> u' x" show "max 0 (ereal (2 * w x)) = liminf (\n. max 0 (ereal (?diff n x)))" unfolding ereal_max_0 proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal) have "(\i. ?diff i x) ----> 2 * w x - \u' x - u' x\" - using u'[OF x] by (safe intro!: tendsto_intros) + using u' by (safe intro!: tendsto_intros) then show "(\i. max 0 (?diff i x)) ----> max 0 (2 * w x)" - by (auto intro!: tendsto_real_max simp add: lim_ereal) + by (auto intro!: tendsto_real_max) qed (rule trivial_limit_sequentially) qed also have "\ \ liminf (\n. \\<^isup>+ x. max 0 (ereal (?diff n x)) \M)" - using u'_borel w u unfolding integrable_def + using borel w u unfolding integrable_def by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF) also have "\ = (\\<^isup>+ x. ereal (2 * w x) \M) - limsup (\n. \\<^isup>+ x. ereal \u n x - u' x\ \M)" @@ -2106,7 +2143,7 @@ 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) + by simp show ?lim proof (rule LIMSEQ_I) @@ -2119,9 +2156,9 @@ proof (safe intro!: exI[of _ N]) fix n assume "N \ n" have "\integral\<^isup>L M (u n) - integral\<^isup>L M u'\ = \(\x. u n x - u' x \M)\" - using u `integrable M u'` by (auto simp: integral_diff) + using u `integrable M u'` by auto also have "\ \ (\x. \u n x - u' x\ \M)" using u `integrable M u'` - by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff) + by (rule_tac integral_triangle_inequality) auto also note N[OF `N \ n`] finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp qed @@ -2139,6 +2176,8 @@ using summable unfolding summable_def by auto from bchoice[OF this] obtain w where w: "\x. x \ space M \ (\i. \f i x\) sums w x" by auto + then have w_borel: "w \ borel_measurable M" unfolding sums_def + by (rule borel_measurable_LIMSEQ) (auto simp: borel[THEN integrableD(1)]) let ?w = "\y. if y \ space M then w y else 0" @@ -2146,13 +2185,16 @@ using sums unfolding summable_def .. have 1: "\n. integrable M (\x. \i = 0.. space M" + have 2: "\j. AE x in M. \\i = 0.. \ ?w x" + using AE_space + proof eventually_elim + fix j x assume [simp]: "x \ space M" have "\\i = 0..< j. f i x\ \ (\i = 0..< j. \f i x\)" by (rule setsum_abs) also have "\ \ w x" using w[of x] series_pos_le[of "\i. \f i x\"] unfolding sums_iff by auto - finally have "\\i = 0.. \ ?w x" by simp } - note 2 = this + finally show "\\i = 0.. \ ?w x" by simp + qed have 3: "integrable M ?w" proof (rule integral_monotone_convergence(1)) @@ -2161,21 +2203,22 @@ have "\n. integrable M (?F n)" using borel by (auto intro!: integral_setsum integrable_abs) thus "\n. integrable M (?w' n)" by (simp cong: integrable_cong) - show "mono ?w'" + show "AE x in M. mono (\n. ?w' n x)" by (auto simp: mono_def le_fun_def intro!: setsum_mono2) - { fix x show "(\n. ?w' n x) ----> ?w x" - using w by (cases "x \ space M") (simp_all add: tendsto_const sums_def) } + show "AE x in M. (\n. ?w' n x) ----> ?w x" + using w by (simp_all add: tendsto_const sums_def) have *: "\n. integral\<^isup>L M (?w' n) = (\i = 0..< n. (\x. \f i x\ \M))" - using borel by (simp add: integral_setsum integrable_abs cong: integral_cong) + using borel by (simp add: integrable_abs cong: integral_cong) from abs_sum show "(\i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def . - qed + qed (simp add: w_borel measurable_If_set) from summable[THEN summable_rabs_cancel] - have 4: "\x. x \ space M \ (\n. \i = 0.. (\i. f i x)" + have 4: "AE x in M. (\n. \i = 0.. (\i. f i x)" by (auto intro: summable_sumr_LIMSEQ_suminf) - note int = integral_dominated_convergence(1,3)[OF 1 2 3 4] + note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 + borel_measurable_suminf[OF integrableD(1)[OF borel]]] from int show "integrable M ?S" by simp @@ -2244,59 +2287,54 @@ 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 (distr M M' T)" "\x. 0 \ f x" + shows "integral\<^isup>P (distr M M' T) f = (\\<^isup>+ x. f (T x) \M)" + using f +proof induct + case (cong f g) + with T show ?case + apply (subst positive_integral_cong[of _ f g]) + apply simp + apply (subst positive_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) + apply (simp add: measurable_def Pi_iff) + apply simp + done +next + case (set A) + then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" + by (auto simp: indicator_def) + from set T show ?case + by (subst positive_integral_cong[OF eq]) + (auto simp add: emeasure_distr intro!: positive_integral_indicator[symmetric] measurable_sets) +qed (simp_all add: measurable_compose[OF T] T positive_integral_cmult positive_integral_add + positive_integral_monotone_convergence_SUP le_fun_def incseq_def) 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 + apply (subst (1 2) positive_integral_max_0[symmetric]) + apply (rule positive_integral_distr') + apply (auto simp: f T) + done 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 + "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^isup>L (distr M M' T) f = (\ x. f (T x) \M)" + unfolding lebesgue_integral_def + by (subst (1 2) positive_integral_distr) auto + +lemma integrable_distr_eq: + assumes T: "T \ measurable M M'" "f \ borel_measurable M'" + shows "integrable (distr M M' T) f \ integrable M (\x. f (T x))" + using T measurable_comp[OF T] + unfolding integrable_def + by (subst (1 2) positive_integral_distr) (auto simp: comp_def) 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 + assumes T: "T \ measurable M M'" shows "integrable (distr M M' T) f \ integrable M (\x. f (T x))" + by (subst integrable_distr_eq[symmetric, OF T]) auto section {* Lebesgue integration on @{const count_space} *} @@ -2329,6 +2367,26 @@ 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_support: + assumes f: "finite {a\A. f a \ 0}" shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" +proof - + have *: "\r::real. 0 < max 0 r \ 0 < r" "\x. max 0 (ereal x) = ereal (max 0 x)" + "\a. a \ A \ 0 < f a \ max 0 (f a) = f a" + "\a. a \ A \ f a < 0 \ max 0 (- f a) = - f a" + "{a \ A. f a \ 0} = {a \ A. 0 < f a} \ {a \ A. f a < 0}" + "({a \ A. 0 < f a} \ {a \ A. f a < 0}) = {}" + by (auto split: split_max) + have "finite {a \ A. 0 < f a}" "finite {a \ A. f a < 0}" + by (auto intro: finite_subset[OF _ f]) + then show ?thesis + unfolding lebesgue_integral_def + apply (subst (1 2) positive_integral_max_0[symmetric]) + apply (subst (1 2) positive_integral_count_space) + apply (auto simp add: * setsum_negf setsum_Un + simp del: ereal_max) + done +qed + 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 @@ -2337,6 +2395,10 @@ apply (auto simp: max_def setsum_subtractf[symmetric] intro!: setsum_cong) done +lemma borel_measurable_count_space[simp, intro!]: + "f \ borel_measurable (count_space A)" + by simp + section {* Measure spaces with an associated density *} definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where @@ -2449,65 +2511,50 @@ (auto elim: eventually_elim2) qed -lemma positive_integral_density: +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 . + assumes g: "g \ borel_measurable M" "\x. 0 \ g x" + shows "integral\<^isup>P (density M f) g = (\\<^isup>+ x. f x * g x \M)" +using g proof induct + case (cong u v) + then show ?case + apply (subst positive_integral_cong[OF cong(3)]) + apply (simp_all cong: positive_integral_cong) + done +next + case (set A) then show ?case + by (simp add: emeasure_density f) +next + case (mult u c) + moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) + ultimately show ?case + by (simp add: f positive_integral_cmult) +next + case (add u v) + moreover then have "\x. f x * (v x + u x) = f x * v x + f x * u x" + by (simp add: ereal_right_distrib) + moreover note f + ultimately show ?case + by (auto simp add: positive_integral_add ereal_zero_le_0_iff intro!: positive_integral_add[symmetric]) +next + case (seq U) + from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" + by eventually_elim (simp add: SUPR_ereal_cmult seq) + from seq f show ?case + apply (simp add: positive_integral_monotone_convergence_SUP) + apply (subst positive_integral_cong_AE[OF eq]) + apply (subst positive_integral_monotone_convergence_SUP_AE) + apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) + done qed +lemma positive_integral_density: + "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ + integral\<^isup>P (density M f) g' = (\\<^isup>+ x. f x * g' x \M)" + by (subst (1 2) positive_integral_max_0[symmetric]) + (auto intro!: positive_integral_cong_AE + simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density') + lemma integral_density: assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" and g: "g \ borel_measurable M" @@ -2615,9 +2662,14 @@ 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)" + "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) +lemma emeasure_point_measure_finite2: + "X \ A \ finite X \ (\i. i \ X \ 0 \ f i) \ emeasure (point_measure A f) X = (\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) diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -9,20 +9,23 @@ imports Finite_Product_Measure begin +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 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 "cube n \ {\\ i. - real n .. \\ i. real n}" +lemma borel_cube[intro]: "cube n \ sets borel" + unfolding cube_def by auto + lemma cube_closed[intro]: "closed (cube n)" unfolding cube_def by auto @@ -154,7 +157,7 @@ then have UN_A[simp, intro]: "\i n. (indicator (\i. A i) :: _ \ real) integrable_on cube n" by (auto simp: sets_lebesgue) show "(\n. ?\ (A n)) = ?\ (\i. A i)" - proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) + 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 @@ -193,7 +196,6 @@ qed qed qed -next qed (auto, fact) lemma has_integral_interval_cube: @@ -279,14 +281,16 @@ lemma lmeasure_finite_has_integral: fixes s :: "'a::ordered_euclidean_space set" - assumes "s \ sets lebesgue" "emeasure lebesgue s = ereal m" "0 \ m" + assumes "s \ sets lebesgue" "emeasure lebesgue s = ereal m" shows "(indicator s has_integral m) UNIV" proof - let ?I = "indicator :: 'a set \ 'a \ real" + have "0 \ m" + using emeasure_nonneg[of lebesgue s] `emeasure lebesgue s = ereal m` by simp have **: "(?I s) integrable_on UNIV \ (\k. integral UNIV (?I (s \ cube k))) ----> integral UNIV (?I s)" proof (intro monotone_convergence_increasing allI ballI) have LIMSEQ: "(\n. integral (cube n) (?I s)) ----> m" - using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1, 3)] . + using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \ m`] . { fix n have "integral (cube n) (?I s) \ m" using cube_subset assms by (intro incseq_le[where L=m] LIMSEQ incseq_def[THEN iffD2] integral_subset_le allI impI) @@ -316,7 +320,7 @@ note ** = conjunctD2[OF this] have m: "m = integral UNIV (?I s)" apply (intro LIMSEQ_unique[OF _ **(2)]) - using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1,3)] integral_indicator_UNIV . + using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \ m`] integral_indicator_UNIV . show ?thesis unfolding m by (intro integrable_integral **) qed @@ -366,14 +370,14 @@ qed lemma has_integral_iff_lmeasure: - "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ emeasure lebesgue A = ereal m)" + "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 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 \ emeasure lebesgue A = ereal m" + show "A \ sets lebesgue \ emeasure lebesgue A = ereal m" by (auto intro: has_integral_nonneg) next - assume "A \ sets lebesgue \ 0 \ m \ emeasure lebesgue A = ereal m" + assume "A \ sets lebesgue \ emeasure lebesgue A = ereal m" then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto qed @@ -450,6 +454,9 @@ by (auto simp: cube_def content_closed_interval_cases setprod_constant) qed simp +lemma lmeasure_complete: "A \ B \ B \ null_sets lebesgue \ A \ null_sets lebesgue" + unfolding negligible_iff_lebesgue_null_sets[symmetric] by (auto simp: negligible_subset) + lemma fixes a b ::"'a::ordered_euclidean_space" shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})" @@ -475,43 +482,44 @@ fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0" using lmeasure_atLeastAtMost[of a a] by simp +lemma AE_lebesgue_singleton: + fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \ a" + by (rule AE_I[where N="{a}"]) auto + declare content_real[simp] lemma fixes a b :: real shows lmeasure_real_greaterThanAtMost[simp]: "emeasure lebesgue {a <.. b} = ereal (if a \ b then b - a else 0)" -proof cases - assume "a < b" - 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"]) +proof - + have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}" + using AE_lebesgue_singleton[of a] + by (intro emeasure_eq_AE) auto then show ?thesis by auto -qed auto +qed lemma fixes a b :: real shows lmeasure_real_atLeastLessThan[simp]: "emeasure lebesgue {a ..< b} = ereal (if a \ b then b - a else 0)" -proof cases - assume "a < b" - 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"]) +proof - + have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}" + using AE_lebesgue_singleton[of b] + by (intro emeasure_eq_AE) auto then show ?thesis by auto -qed auto +qed lemma fixes a b :: real shows lmeasure_real_greaterThanLessThan[simp]: "emeasure lebesgue {a <..< b} = ereal (if a \ b then b - a else 0)" -proof cases - assume "a < b" - 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"]) +proof - + have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}" + using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b] + by (intro emeasure_eq_AE) auto then show ?thesis by auto -qed auto +qed subsection {* Lebesgue-Borel measure *} @@ -544,6 +552,61 @@ by (intro exI[of _ A]) (auto simp: subset_eq) 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 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 + { 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 lebesgue_real_affine: + 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 simp + +lemma lebesgue_integral_real_affine: + fixes c :: real assumes c: "c \ 0" and f: "f \ borel_measurable borel" + shows "(\ x. f x \ lborel) = \c\ * (\ x. f (t + c * x) \lborel)" + by (subst lebesgue_real_affine[OF c, of t]) + (simp add: f integral_density integral_distr lebesgue_integral_cmult) + subsection {* Lebesgue integrable implies Gauge integrable *} lemma has_integral_cmult_real: @@ -632,6 +695,69 @@ qed qed +lemma borel_measurable_real_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: + fixes u :: "'a \ real" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" + assumes seq: "\U u. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\x. (\i. U i x) ----> u x) \ (\i. P (U i)) \ incseq U \ P u" + shows "P u" +proof - + have "(\x. ereal (u x)) \ borel_measurable M" + using u by auto + then obtain U where U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and + "\x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\i x. 0 \ U i x" + using borel_measurable_implies_simple_function_sequence'[of u M] by auto + then have u_eq: "\x. ereal (u x) = (SUP i. U i x)" + using u by (auto simp: max_def) + + have [simp]: "\i x. U i x \ \" using U by (auto simp: image_def eq_commute) + + { fix i x have [simp]: "U i x \ -\" using nn[of i x] by auto } + note this[simp] + + show "P u" + proof (rule seq) + show "\i. (\x. real (U i x)) \ borel_measurable M" + using U by (auto intro: borel_measurable_simple_function) + show "\i x. 0 \ real (U i x)" + using nn by (auto simp: real_of_ereal_pos) + show "incseq (\i x. real (U i x))" + using U(2) by (auto simp: incseq_def image_iff le_fun_def intro!: real_of_ereal_positive_mono nn) + then show "\x. (\i. real (U i x)) ----> u x" + by (intro SUP_eq_LIMSEQ[THEN iffD1]) + (auto simp: incseq_mono incseq_def le_fun_def u_eq ereal_real + intro!: arg_cong2[where f=SUPR] ext) + show "\i. P (\x. real (U i x))" + proof (rule cong) + fix x i assume x: "x \ space M" + have [simp]: "\A x. real (indicator A x :: ereal) = indicator A x" + by (auto simp: indicator_def one_ereal_def) + { fix y assume "y \ U i ` space M" + then have "0 \ y" "y \ \" using nn by auto + then have "\y * indicator (U i -` {y} \ space M) x\ \ \" + by (auto simp: indicator_def) } + then show "real (U i x) = (\y \ U i ` space M. real y * indicator (U i -` {y} \ space M) x)" + unfolding simple_function_indicator_representation[OF U(1) x] + by (subst setsum_real_of_ereal[symmetric]) auto + next + fix i + have "finite (U i ` space M)" "\x. x \ U i ` space M \ 0 \ x""\x. x \ U i ` space M \ x \ \" + using U(1) nn by (auto simp: simple_function_def) + then show "P (\x. \y \ U i ` space M. real y * indicator (U i -` {y} \ space M) x)" + proof induct + case empty then show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + qed (auto intro!: add mult set simple_functionD U setsum_nonneg borel_measurable_setsum mult_nonneg_nonneg real_of_ereal_pos) + qed (auto intro: borel_measurable_simple_function U simple_functionD intro!: borel_measurable_setsum borel_measurable_times) + qed +qed + +lemma ereal_indicator: "ereal (indicator A x) = indicator A x" + by (auto simp: indicator_def one_ereal_def) + lemma positive_integral_has_integral: fixes f :: "'a::ordered_euclidean_space \ ereal" assumes f: "f \ borel_measurable lebesgue" "range f \ {0..<\}" "integral\<^isup>P lebesgue f \ \" @@ -777,201 +903,22 @@ unfolding lebesgue_integral_eq_borel[OF borel] by simp qed -subsection {* Equivalence between product spaces and euclidean spaces *} - -definition e2p :: "'a::ordered_euclidean_space \ (nat \ real)" where - "e2p x = (\i\{.. real) \ 'a::ordered_euclidean_space" where - "p2e x = (\\ i. x i)" - -lemma e2p_p2e[simp]: - "x \ extensional {.. e2p (p2e x::'a::ordered_euclidean_space) = x" - by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def) - -lemma p2e_e2p[simp]: - "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" - by default - -interpretation lborel_space: finite_product_sigma_finite "\x. lborel::real measure" "{..x\A. \y. P x y) \ (\f. \x\A. P x (f x))" - by metis - -lemma sets_product_borel: - assumes I: "finite I" - shows "sets (\\<^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 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 ?P _") -proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2]) - fix x i - let ?A = "{w \ space ?P. (p2e w :: 'a) $$ 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 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: - "(lborel :: 'a measure) = distr (\\<^isub>M i\{.. space ?P = (\\<^isub>E i\{.. ereal" - assumes f: "f \ borel_measurable borel" - shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(\\<^isub>M i\{.. real" - shows "integrable lborel f \ - integrable (\\<^isub>M i\{..x. f (p2e x))" - (is "_ \ integrable ?B ?f") -proof - assume "integrable lborel f" - moreover then have f: "f \ borel_measurable borel" - by auto - moreover with measurable_p2e - have "f \ p2e \ borel_measurable ?B" - by (rule measurable_comp) - ultimately show "integrable ?B ?f" - 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 measure)" - by (auto intro!: measurable_e2p) - then have "f \ borel_measurable borel" - by (simp cong: measurable_cong) - ultimately show "integrable lborel f" - by (simp add: borel_fubini_positiv_integral integrable_def) -qed - -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) \((\\<^isub>M i\{.. 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 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 simp - -lemma borel_cube[intro]: "cube n \ sets borel" - unfolding cube_def by auto - lemma integrable_on_cmult_iff: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on s \ f integrable_on s" using integrable_cmul[of "\x. c * f x" s "1 / c"] integrable_cmul[of f s c] `c \ 0` by auto -lemma positive_integral_borel_has_integral: +lemma positive_integral_lebesgue_has_integral: fixes f :: "'a::ordered_euclidean_space \ real" - assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" + assumes f_borel: "f \ borel_measurable lebesgue" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" - shows "(\\<^isup>+x. f x \lborel) = I" + shows "(\\<^isup>+x. f x \lebesgue) = I" proof - - from f_borel have "(\x. ereal (f x)) \ borel_measurable borel" by auto + from f_borel have "(\x. ereal (f x)) \ borel_measurable lebesgue" by auto from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this - have lebesgue_eq: "(\\<^isup>+ x. ereal (f x) \lebesgue) = (\\<^isup>+ x. ereal (f x) \lborel)" - using f_borel by (intro lebesgue_positive_integral_eq_borel) auto - also have "\ = (SUP i. integral\<^isup>S lborel (F i))" + have "(\\<^isup>+ x. ereal (f x) \lebesgue) = (SUP i. integral\<^isup>S lebesgue (F i))" using F by (subst positive_integral_monotone_convergence_simple) (simp_all add: positive_integral_max_0 simple_function_def) @@ -1043,11 +990,8 @@ unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast moreover have "0 \ integral\<^isup>S lebesgue (F i)" using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def) - moreover have "integral\<^isup>S lebesgue (F i) = integral\<^isup>S lborel (F i)" - using F(1)[of i, THEN borel_measurable_simple_function] - by (rule lebesgue_simple_integral_eq_borel) - ultimately show "integral\<^isup>S lborel (F i) \ ereal I" - by (cases "integral\<^isup>S lborel (F i)") auto + ultimately show "integral\<^isup>S lebesgue (F i) \ ereal I" + by (cases "integral\<^isup>S lebesgue (F i)") auto qed also have "\ < \" by simp finally have finite: "(\\<^isup>+ x. ereal (f x) \lebesgue) \ \" by simp @@ -1059,14 +1003,142 @@ with I have "I = real (\\<^isup>+ x. ereal (f x) \lebesgue)" by (rule has_integral_unique) with finite positive_integral_positive[of _ "\x. ereal (f x)"] show ?thesis - by (cases "\\<^isup>+ x. ereal (f x) \lborel") (auto simp: lebesgue_eq) + by (cases "\\<^isup>+ x. ereal (f x) \lebesgue") auto qed -lemma has_integral_iff_positive_integral: +lemma has_integral_iff_positive_integral_lebesgue: + fixes f :: "'a::ordered_euclidean_space \ real" + assumes f: "f \ borel_measurable lebesgue" "\x. 0 \ f x" + shows "(f has_integral I) UNIV \ integral\<^isup>P lebesgue f = I" + using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f] + by (auto simp: subset_eq) + +lemma has_integral_iff_positive_integral_lborel: fixes f :: "'a::ordered_euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" shows "(f has_integral I) UNIV \ integral\<^isup>P lborel f = I" - using f positive_integral_borel_has_integral[of f I] positive_integral_has_integral[of f] - by (auto simp: subset_eq borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel) + using assms + by (subst has_integral_iff_positive_integral_lebesgue) + (auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel) + +subsection {* Equivalence between product spaces and euclidean spaces *} + +definition e2p :: "'a::ordered_euclidean_space \ (nat \ real)" where + "e2p x = (\i\{.. real) \ 'a::ordered_euclidean_space" where + "p2e x = (\\ i. x i)" + +lemma e2p_p2e[simp]: + "x \ extensional {.. e2p (p2e x::'a::ordered_euclidean_space) = x" + by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def) + +lemma p2e_e2p[simp]: + "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" + by default + +interpretation lborel_space: finite_product_sigma_finite "\x. lborel::real measure" "{..x\A. \y. P x y) \ (\f. \x\A. P x (f x))" + by metis + +lemma sets_product_borel: + assumes I: "finite I" + shows "sets (\\<^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 reals_Archimedean2) + +lemma measurable_e2p: + "e2p \ 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 ?P _") +proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2]) + fix x i + let ?A = "{w \ space ?P. (p2e w :: 'a) $$ i \ x}" + assume "i < DIM('a)" + then have "?A = (\\<^isub>E j\{.. sets ?P" + by auto +qed + +lemma lborel_eq_lborel_space: + "(lborel :: 'a measure) = distr (\\<^isub>M i\{.. space ?P = (\\<^isub>E i\{.. ereal" + assumes f: "f \ borel_measurable borel" + shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(\\<^isub>M i\{.. real" + shows "integrable lborel f \ + integrable (\\<^isub>M i\{..x. f (p2e x))" + (is "_ \ integrable ?B ?f") +proof + assume "integrable lborel f" + moreover then have f: "f \ borel_measurable borel" + by auto + moreover with measurable_p2e + have "f \ p2e \ borel_measurable ?B" + by (rule measurable_comp) + ultimately show "integrable ?B ?f" + 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 measure)" + by (auto intro!: measurable_e2p) + then have "f \ borel_measurable borel" + by (simp cong: measurable_cong) + ultimately show "integrable lborel f" + by (simp add: borel_fubini_positiv_integral integrable_def) +qed + +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) \((\\<^isub>M i\{.. (\n. (\i\n. f i)) ----> x" + unfolding sums_def + apply (subst LIMSEQ_Suc_iff[symmetric]) + unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .. + lemma suminf_cmult_indicator: fixes f :: "nat \ ereal" assumes "disjoint_family A" "x \ A i" "\i. 0 \ f i" @@ -90,6 +96,9 @@ definition increasing where "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" +lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) +lemma positiveD2: "positive M f \ A \ M \ 0 \ f A" by (auto simp: positive_def) + lemma positiveD_empty: "positive M f \ f {} = 0" by (auto simp add: positive_def) @@ -240,6 +249,143 @@ finally show "(\i. \ (F i)) = \ (\i. F i)" . qed +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 \ 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 \ 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) + moreover have "(\n. (\i=0.. (\i. f (disjointed A i))" + using f(1)[unfolded positive_def] dA + by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos) + from LIMSEQ_Suc[OF this] + have "(\n. (\i\n. f (disjointed A i))) ----> (\i. f (disjointed A i))" + unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost . + moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" + 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 \ 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) \ 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))" + using A + by (intro additive_sum[OF f, of _ A, symmetric]) + (auto intro: disjoint_family_on_mono[where B=UNIV]) + ultimately + have "(\i. f (A i)) sums f (\i. A i)" + unfolding sums_def2 by simp + from sums_unique[OF this] + show "(\i. f (A i)) = f (\i. A i)" by simp +qed + +lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: + assumes f: "positive M f" "additive M f" + 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 \ 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 \ 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 \ 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))" + using A by (auto simp: decseq_def intro!: f_mono) + have decseq: "decseq (\i. A i - (\i. A i))" + using A by (auto simp: decseq_def) + then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" + using A unfolding decseq_def by (auto intro!: f_mono Diff) + have "f (\x. A x) \ f (A 0)" + using A by (auto intro!: f_mono) + then have f_Int_fin: "f (\x. A x) \ \" + using A by auto + { fix i + have "f (A i - (\i. A i)) \ f (A i)" using A by (auto intro!: f_mono) + then have "f (A i - (\i. A i)) \ \" + using A by auto } + 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)) \ M" "(\i. A i - (\i. A i)) = {}" + using A by auto + qed + from INF_Lim_ereal[OF decseq_f this] + have "(INF n. f (A n - (\i. A i))) = 0" . + moreover have "(INF n. f (\i. A i)) = f (\i. A i)" + by auto + ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" + using A(4) f_fin f_Int_fin + by (subst INFI_ereal_add) (auto simp: decseq_f) + moreover { + fix n + have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" + using A by (subst f(2)[THEN additiveD]) auto + also have "(A n - (\i. A i)) \ (\i. A i) = A n" + by auto + finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } + ultimately have "(INF n. f (A n)) = f (\i. A i)" + by simp + with LIMSEQ_ereal_INFI[OF decseq_fA] + show "(\i. f (A i)) ----> f (\i. A i)" by simp +qed + +lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: + 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\M. \x. f A = ereal x" + proof + 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] + from A have "(\i. f ((\i. A i) - A i)) ----> 0" + by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) + moreover + { fix i + have "f ((\i. A i) - A i) + f (A i) = f ((\i. A i) - A i \ A i)" + using A by (intro f(2)[THEN additiveD, symmetric]) auto + also have "(\i. A i) - A i \ A i = (\i. A i)" + by auto + finally have "f' (\i. A i) - f' (A i) = f' ((\i. A i) - A i)" + using A by (subst (asm) (1 2 3) f') auto + then have "f ((\i. A i) - A i) = ereal (f' (\i. A i) - f' (A i))" + using A f' by auto } + ultimately have "(\i. f' (\i. A i) - f' (A i)) ----> 0" + by (simp add: zero_ereal_def) + then have "(\i. f' (A i)) ----> f' (\i. A i)" + by (rule LIMSEQ_diff_approach_zero2[OF tendsto_const]) + then show "(\i. f (A i)) ----> f (\i. A i)" + using A by (subst (1 2) f') auto +qed + +lemma (in ring_of_sets) empty_continuous_imp_countably_additive: + 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 + section {* Properties of @{const emeasure} *} lemma emeasure_positive: "positive (sets M) (emeasure M)" @@ -314,84 +460,21 @@ 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 Lim_emeasure_incseq: + "range A \ sets M \ incseq A \ (\i. (emeasure M (A i))) ----> emeasure M (\i. A i)" + using emeasure_countably_additive + by (auto simp add: countably_additive_iff_continuous_from_below emeasure_positive emeasure_additive) 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: +lemma SUP_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 + shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" + using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] + by (simp add: LIMSEQ_unique) lemma decseq_emeasure: assumes "range B \ sets M" "decseq B" @@ -545,70 +628,78 @@ 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) \ \" + and A: "range A \ E" "(\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)}" + let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact - { fix F assume "F \ E" and "emeasure M F \ \" + have "space M = \" + using top[of M] space_closed[of M] S.top S.space_closed `sets M = sigma_sets \ E` by blast + + { fix F D assume "F \ E" and "?\ 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 + have "?\ F \ \" using `?\ F \ \` `F \ E` eq by simp + assume "D \ sets M" + with `Int_stable E` `E \ Pow \` have "emeasure M (F \ D) = emeasure N (F \ D)" + unfolding M + proof (induct rule: sigma_sets_induct_disjoint) + case (basic A) + then have "F \ A \ E" using `Int_stable E` `F \ E` by (auto simp: Int_stable_def) + then show ?case using eq 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) + case empty then show ?case by simp next - fix A assume *: "A \ sigma_sets \ E \ emeasure M (F \ A) = emeasure N (F \ A)" + case (compl 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 \ E` S.sets_into_space by (auto simp: M) + have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) + then have "?\ (F \ A) \ \" using `?\ F \ \` by auto + have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) + then have "?\ (F \ A) \ \" using `?\ F \ \` by auto + then have "?\ (F \ (\ - A)) = ?\ F - ?\ (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) \ \` + also have "\ = ?\ F - ?\ (F \ A)" using eq `F \ E` compl by simp + also have "\ = ?\ (F \ (\ - A))" unfolding ** + using `F \ A \ sigma_sets \ E` `?\ (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 + finally show ?case + using `space M = \` 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 } + case (union A) + then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" + by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) + with A show ?case + by auto + qed } 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) + have [simp, intro]: "\i. A i \ sets M" + using A(1) by (auto simp: subset_eq M) + fix F assume "F \ sets M" + let ?D = "disjointed (\i. F \ A i)" + from `space M = \` have F_eq: "F = (\i. ?D i)" + using `F \ sets M`[THEN sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) + have [simp, intro]: "\i. ?D i \ sets M" + using range_disjointed_sets[of "\i. F \ A i" M] `F \ sets M` + by (auto simp: subset_eq) + have "disjoint_family ?D" + by (auto simp: disjoint_family_disjointed) + moreover + { fix i + have "A i \ ?D i = ?D i" + by (auto simp: disjointed_def) + then have "emeasure M (?D i) = emeasure N (?D i)" + using *[of "A i" "?D i", OF _ A(3)] A(1) by auto } + ultimately show "emeasure M F = emeasure N F" + using N M + apply (subst (1 2) F_eq) + apply (subst (1 2) suminf_emeasure[symmetric]) + apply auto + done qed qed @@ -1016,6 +1107,24 @@ (auto simp: emeasure_distr measurable_def) qed +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 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) @@ -1295,103 +1404,81 @@ unfolding measurable_def by auto qed +lemma strict_monoI_Suc: + assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" + unfolding strict_mono_def +proof safe + fix n m :: nat assume "n < m" then show "f n < f m" + by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) +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 "X \ Pow A" using `X \ A` by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) - - show "positive (Pow A) ?M" + show positive: "positive (Pow A) ?M" by (auto simp: positive_def) + have additive: "additive (Pow A) ?M" + by (auto simp: card_Un_disjoint additive_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 + interpret ring_of_sets A "Pow A" + by (rule ring_of_setsI) auto + show "countably_additive (Pow A) ?M" + unfolding countably_additive_iff_continuous_from_below[OF positive additive] + proof safe + fix F :: "nat \ 'a set" assume "incseq F" + show "(\i. ?M (F i)) ----> ?M (\i. F i)" + proof cases + assume "\i. \j\i. F i = F j" + then guess i .. note i = this + { fix j from i `incseq F` have "F j \ F i" + by (cases "i \ j") (auto simp: incseq_def) } + then have eq: "(\i. F i) = F i" + by auto + with i show ?thesis + by (auto intro!: Lim_eventually eventually_sequentiallyI[where c=i]) + next + assume "\ (\i. \j\i. F i = F j)" + then obtain f where "\i. i \ f i" "\i. F i \ F (f i)" by metis + moreover then have "\i. F i \ F (f i)" using `incseq F` by (auto simp: incseq_def) + ultimately have *: "\i. F i \ F (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_finite) auto - 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]) + have "incseq (\i. ?M (F i))" + using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset) + then have "(\i. ?M (F i)) ----> (SUP n. ?M (F n))" + by (rule LIMSEQ_ereal_SUPR) - 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 + moreover have "(SUP n. ?M (F n)) = \" + proof (rule SUP_PInfty) + fix n :: nat show "\k::nat\UNIV. ereal n \ ?M (F k)" + proof (induct n) + case (Suc n) + then guess k .. note k = this + moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" + using `F k \ F (f k)` by (simp add: psubset_card_mono) + moreover have "finite (F (f k)) \ finite (F k)" + using `k \ f k` `incseq F` by (auto simp: incseq_def dest: finite_subset) + ultimately show ?case + by (auto intro!: exI[of _ "f k"]) + qed auto qed + + moreover + have "inj (\n. F ((f ^^ n) 0))" + by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) + then have 1: "infinite (range (\i. F ((f ^^ i) 0)))" + by (rule range_inj_infinite) + have "infinite (Pow (\i. F i))" + by (rule infinite_super[OF _ 1]) auto + then have "infinite (\i. F i)" + by auto + + ultimately show ?thesis by auto + qed qed - show "X \ Pow A" using `X \ A` by simp qed lemma emeasure_count_space_finite[simp]: diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -219,26 +219,26 @@ with AE_in_set_eq_1 assms show ?thesis by simp qed -lemma (in prob_space) prob_space_increasing: "increasing M prob" +lemma (in finite_measure) prob_space_increasing: "increasing M (measure M)" by (auto intro!: finite_measure_mono simp: increasing_def) -lemma (in prob_space) prob_zero_union: - assumes "s \ events" "t \ events" "prob t = 0" - shows "prob (s \ t) = prob s" +lemma (in finite_measure) prob_zero_union: + assumes "s \ sets M" "t \ sets M" "measure M t = 0" + shows "measure M (s \ t) = measure M s" using assms proof - - have "prob (s \ t) \ prob s" + have "measure M (s \ t) \ measure M s" using finite_measure_subadditive[of s t] assms by auto - moreover have "prob (s \ t) \ prob s" + moreover have "measure M (s \ t) \ measure M s" using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed -lemma (in prob_space) prob_eq_compl: - assumes "s \ events" "t \ events" - assumes "prob (space M - s) = prob (space M - t)" - shows "prob s = prob t" - using assms prob_compl by auto +lemma (in finite_measure) prob_eq_compl: + assumes "s \ sets M" "t \ sets M" + assumes "measure M (space M - s) = measure M (space M - t)" + shows "measure M s = measure M t" + using assms finite_measure_compl by auto lemma (in prob_space) prob_one_inter: assumes events:"s \ events" "t \ events" @@ -253,26 +253,26 @@ using events by (auto intro!: prob_eq_compl[of "s \ t" s]) qed -lemma (in prob_space) prob_eq_bigunion_image: - assumes "range f \ events" "range g \ events" +lemma (in finite_measure) prob_eq_bigunion_image: + assumes "range f \ sets M" "range g \ sets M" assumes "disjoint_family f" "disjoint_family g" - assumes "\ n :: nat. prob (f n) = prob (g n)" - shows "(prob (\ i. f i) = prob (\ i. g i))" + assumes "\ n :: nat. measure M (f n) = measure M (g n)" + shows "measure M (\ i. f i) = measure M (\ i. g i)" using assms proof - - have a: "(\ i. prob (f i)) sums (prob (\ i. f i))" + have a: "(\ i. measure M (f i)) sums (measure M (\ i. f i))" by (rule finite_measure_UNION[OF assms(1,3)]) - have b: "(\ i. prob (g i)) sums (prob (\ i. g i))" + have b: "(\ i. measure M (g i)) sums (measure M (\ i. g i))" by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed -lemma (in prob_space) prob_countably_zero: - assumes "range c \ events" - assumes "\ i. prob (c i) = 0" - shows "prob (\ i :: nat. c i) = 0" +lemma (in finite_measure) prob_countably_zero: + assumes "range c \ sets M" + assumes "\ i. measure M (c i) = 0" + shows "measure M (\ i :: nat. c i) = 0" proof (rule antisym) - show "prob (\ i :: nat. c i) \ 0" + show "measure M (\ i :: nat. c i) \ 0" using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) qed (simp add: measure_nonneg) @@ -316,7 +316,7 @@ lemma (in prob_space) expectation_less: assumes [simp]: "integrable M X" - assumes gt: "\x\space M. X x < b" + assumes gt: "AE x in M. X x < b" shows "expectation X < b" proof - have "expectation X < expectation (\x. b)" @@ -327,7 +327,7 @@ lemma (in prob_space) expectation_greater: assumes [simp]: "integrable M X" - assumes gt: "\x\space M. a < X x" + assumes gt: "AE x in M. a < X x" shows "a < expectation X" proof - have "expectation (\x. a) < expectation X" @@ -338,13 +338,13 @@ lemma (in prob_space) jensens_inequality: fixes a b :: real - assumes X: "integrable M X" "X ` space M \ I" + assumes X: "integrable M X" "AE x in M. X x \ I" assumes I: "I = {a <..< b} \ I = {a <..} \ I = {..< b} \ I = UNIV" assumes q: "integrable M (\x. q (X x))" "convex_on I q" shows "q (expectation X) \ expectation (\x. q (X x))" proof - let ?F = "\x. Inf ((\t. (q x - q t) / (x - t)) ` ({x<..} \ I))" - from not_empty X(2) have "I \ {}" by auto + from X(2) AE_False have "I \ {}" by auto from I have "open I" by auto @@ -376,9 +376,12 @@ 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 - lebesgue_integral_const X q convex_le_Inf_differential) - (auto simp: interior_open) + apply (intro integral_mono_AE integral_add integral_cmult integral_diff + lebesgue_integral_const X q) + apply (elim eventually_elim1) + apply (intro convex_le_Inf_differential) + apply (auto simp: interior_open q) + done finally show "k \ expectation (\w. q (X w))" using x by auto qed finally show "q (expectation X) \ expectation (\x. q (X x))" . @@ -407,7 +410,7 @@ sublocale pair_prob_space \ P: prob_space "M1 \\<^isub>M M2" proof 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) + by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure) qed locale product_prob_space = product_sigma_finite M for M :: "'i \ 'a measure" + @@ -503,6 +506,16 @@ 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_emeasure: + "distributed M N X f \ A \ sets N \ emeasure M (X -` A \ space M) = (\\<^isup>+x. f x * indicator A x \N)" + by (auto simp: distributed_measurable distributed_AE distributed_borel_measurable + distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr) + +lemma distributed_positive_integral: + "distributed M N X f \ g \ borel_measurable N \ (\\<^isup>+x. f x * g x \N) = (\\<^isup>+x. g (X x) \M)" + by (auto simp: distributed_measurable distributed_AE distributed_borel_measurable + distributed_distr_eq_density[symmetric] positive_integral_density[symmetric] positive_integral_distr) + 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 @@ -523,54 +536,116 @@ finally show ?thesis . qed -lemma distributed_marginal_eq_joint: - assumes T: "sigma_finite_measure T" - assumes S: "sigma_finite_measure S" +lemma (in prob_space) distributed_unique: 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) + assumes Py: "distributed M S X Py" + shows "AE x in S. Px x = Py x" +proof - + interpret X: prob_space "distr M S X" + using distributed_measurable[OF Px] by (rule prob_space_distr) + have "sigma_finite_measure (distr M S X)" .. + with sigma_finite_density_unique[of Px S Py ] Px Py + show ?thesis + by (auto simp: distributed_def) +qed + +lemma (in prob_space) distributed_jointI: + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes X[simp]: "X \ measurable M S" and Y[simp]: "Y \ measurable M T" + assumes f[simp]: "f \ borel_measurable (S \\<^isub>M T)" "AE x in S \\<^isub>M T. 0 \ f x" + assumes eq: "\A B. A \ sets S \ B \ sets T \ + emeasure M {x \ space M. X x \ A \ Y x \ B} = (\\<^isup>+x. (\\<^isup>+y. f (x, y) * indicator B y \T) * indicator A x \S)" + shows "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) f" + 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 - 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 + from ST.sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('b \ 'c) set" .. note F = this + let ?E = "{a \ b |a b. a \ sets S \ b \ sets T}" + let ?P = "S \\<^isub>M T" + show "distr M ?P (\x. (X x, Y x)) = density ?P f" (is "?L = ?R") + proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]]) + show "?E \ Pow (space ?P)" + using space_closed[of S] space_closed[of T] by (auto simp: space_pair_measure) + show "sets ?L = sigma_sets (space ?P) ?E" + by (simp add: sets_pair_measure space_pair_measure) + then show "sets ?R = sigma_sets (space ?P) ?E" + by simp + next + interpret L: prob_space ?L + by (rule prob_space_distr) (auto intro!: measurable_Pair) + show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?L (F i) \ \" + using F by (auto simp: space_pair_measure) + next + fix E assume "E \ ?E" + then obtain A B where E[simp]: "E = A \ B" and A[simp]: "A \ sets S" and B[simp]: "B \ sets T" by auto + have "emeasure ?L E = emeasure M {x \ space M. X x \ A \ Y x \ B}" + by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair) + also have "\ = (\\<^isup>+x. (\\<^isup>+y. (f (x, y) * indicator B y) * indicator A x \T) \S)" + by (auto simp add: eq measurable_Pair measurable_compose[OF _ f(1)] positive_integral_multc + intro!: positive_integral_cong) + also have "\ = emeasure ?R E" + by (auto simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric] + intro!: positive_integral_cong split: split_indicator) + finally show "emeasure ?L E = emeasure ?R E" . + qed +qed (auto intro!: measurable_Pair) + +lemma (in prob_space) distributed_swap: + assumes "sigma_finite_measure S" "sigma_finite_measure T" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "distributed M (T \\<^isub>M S) (\x. (Y x, X x)) (\(x, y). Pxy (y, x))" +proof - + 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 + interpret TS: pair_sigma_finite T S by default + + note measurable_Pxy = measurable_compose[OF _ distributed_borel_measurable[OF Pxy]] + show ?thesis + apply (subst TS.distr_pair_swap) + unfolding distributed_def + proof safe + let ?D = "distr (S \\<^isub>M T) (T \\<^isub>M S) (\(x, y). (y, x))" + show 1: "(\(x, y). Pxy (y, x)) \ borel_measurable ?D" + by (auto simp: measurable_split_conv intro!: measurable_Pair measurable_Pxy) + with Pxy + show "AE x in distr (S \\<^isub>M T) (T \\<^isub>M S) (\(x, y). (y, x)). 0 \ (case x of (x, y) \ Pxy (y, x))" + by (subst AE_distr_iff) + (auto dest!: distributed_AE + simp: measurable_split_conv split_beta + intro!: measurable_Pair borel_measurable_ereal_le) + show 2: "random_variable (distr (S \\<^isub>M T) (T \\<^isub>M S) (\(x, y). (y, x))) (\x. (Y x, X x))" + using measurable_compose[OF distributed_measurable[OF Pxy] measurable_fst] + using measurable_compose[OF distributed_measurable[OF Pxy] measurable_snd] + by (auto intro!: measurable_Pair) + { fix A assume A: "A \ sets (T \\<^isub>M S)" + let ?B = "(\(x, y). (y, x)) -` A \ space (S \\<^isub>M T)" + from sets_into_space[OF A] + have "emeasure M ((\x. (Y x, X x)) -` A \ space M) = + emeasure M ((\x. (X x, Y x)) -` ?B \ space M)" + by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure) + also have "\ = (\\<^isup>+ x. Pxy x * indicator ?B x \(S \\<^isub>M T))" + using Pxy A by (intro distributed_emeasure measurable_sets) (auto simp: measurable_split_conv measurable_Pair) + finally have "emeasure M ((\x. (Y x, X x)) -` A \ space M) = + (\\<^isup>+ x. Pxy x * indicator A (snd x, fst x) \(S \\<^isub>M T))" + by (auto intro!: positive_integral_cong split: split_indicator) } + note * = this + show "distr M ?D (\x. (Y x, X x)) = density ?D (\(x, y). Pxy (y, x))" + apply (intro measure_eqI) + apply (simp_all add: emeasure_distr[OF 2] emeasure_density[OF 1]) + apply (subst positive_integral_distr) + apply (auto intro!: measurable_pair measurable_Pxy * simp: comp_def split_beta) + done + qed qed 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)" + defines "Px \ \x. (\\<^isup>+z. Pxy (x, z) \T)" shows "distributed M S X Px" unfolding distributed_def proof safe @@ -585,18 +660,15 @@ 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) + from Pxy show borel: "Px \ borel_measurable S" + by (auto intro!: ST.positive_integral_fst_measurable dest!: distributed_borel_measurable simp: Px_def) 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))" + have "(\\<^isup>+ x. max 0 (- 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) - + by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_borel_measurable distributed_AE) + show "distr M S X = density S Px" proof (rule measure_eqI) fix A assume A: "A \ sets (distr M S X)" @@ -605,31 +677,100 @@ 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" + also have "\ = \\<^isup>+ x. \\<^isup>+ y. 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" + also have "\ = \\<^isup>+ x. 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 + using Pxy[THEN distributed_AE, THEN ST.AE_pair] 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))" + fix x assume "x \ space S" "AE y in T. 0 \ 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" . + ultimately have "(\\<^isup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T) = (\\<^isup>+ y. Pxy (x, y) \T) * indicator A x" + using Pxy[THEN distributed_borel_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong) + also have "(\\<^isup>+ y. Pxy (x, y) \T) = Px x" + by (simp add: Px_def ereal_real positive_integral_positive) + finally show "(\\<^isup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T) = 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)" + show "AE x in S. 0 \ Px x" by (simp add: Px_def positive_integral_positive real_of_ereal_pos) qed +lemma (in prob_space) distr_marginal2: + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "distributed M T Y (\y. (\\<^isup>+x. Pxy (x, y) \S))" + using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp + +lemma (in prob_space) distributed_marginal_eq_joint1: + assumes T: "sigma_finite_measure T" + assumes S: "sigma_finite_measure S" + assumes Px: "distributed M S X Px" + assumes Pxy: "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) Pxy" + shows "AE x in S. Px x = (\\<^isup>+y. Pxy (x, y) \T)" + using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique) + +lemma (in prob_space) distributed_marginal_eq_joint2: + assumes T: "sigma_finite_measure T" + assumes S: "sigma_finite_measure S" + 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)" + using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique) + +lemma (in prob_space) distributed_joint_indep': + assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" + assumes X: "distributed M S X Px" and Y: "distributed M T Y Py" + assumes indep: "distr M S X \\<^isub>M distr M T Y = distr M (S \\<^isub>M T) (\x. (X x, Y x))" + shows "distributed M (S \\<^isub>M T) (\x. (X x, Y x)) (\(x, y). Px x * Py y)" + 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 + + interpret X: prob_space "density S Px" + unfolding distributed_distr_eq_density[OF X, symmetric] + using distributed_measurable[OF X] + by (rule prob_space_distr) + have sf_X: "sigma_finite_measure (density S Px)" .. + + interpret Y: prob_space "density T Py" + unfolding distributed_distr_eq_density[OF Y, symmetric] + using distributed_measurable[OF Y] + by (rule prob_space_distr) + have sf_Y: "sigma_finite_measure (density T Py)" .. + + show "distr M (S \\<^isub>M T) (\x. (X x, Y x)) = density (S \\<^isub>M T) (\(x, y). Px x * Py y)" + unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y] + using distributed_borel_measurable[OF X] distributed_AE[OF X] + using distributed_borel_measurable[OF Y] distributed_AE[OF Y] + by (rule pair_measure_density[OF _ _ _ _ S T sf_X sf_Y]) + + show "random_variable (S \\<^isub>M T) (\x. (X x, Y x))" + using distributed_measurable[OF X] distributed_measurable[OF Y] + by (auto intro: measurable_Pair) + + show Pxy: "(\(x, y). Px x * Py y) \ borel_measurable (S \\<^isub>M T)" + by (auto simp: split_beta' + intro!: measurable_compose[OF _ distributed_borel_measurable[OF X]] + measurable_compose[OF _ distributed_borel_measurable[OF Y]]) + + show "AE x in S \\<^isub>M T. 0 \ (case x of (x, y) \ Px x * Py y)" + apply (intro ST.AE_pair_measure borel_measurable_ereal_le Pxy borel_measurable_const) + using distributed_AE[OF X] + apply eventually_elim + using distributed_AE[OF Y] + apply eventually_elim + apply auto + done +qed + definition "simple_distributed M X f \ distributed M (count_space (X`space M)) X (\x. ereal (f x)) \ finite (X`space M)" @@ -792,19 +933,21 @@ 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], + from distributed_marginal_eq_joint2[OF + sigma_finite_measure_count_space_finite + sigma_finite_measure_count_space_finite + 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] + 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 prob_space_uniform_measure: assumes A: "emeasure M A \ 0" "emeasure M A \ \" shows "prob_space (uniform_measure M A)" diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Wed Oct 10 13:30:50 2012 +0200 @@ -776,15 +776,16 @@ assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" 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 in M. f x = g x)" - (is "(\A\sets M. ?P f A = ?P g A) \ _") + shows "density M f = density M g \ (AE x in M. f x = g x)" proof (intro iffI ballI) 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) + with borel show "density M f = density M g" + by (auto intro: density_cong) next - assume eq: "\A\sets M. ?P f A = ?P g A" + let ?P = "\f A. \\<^isup>+ x. f x * indicator A x \M" + assume "density M f = density M g" + with borel have eq: "\A\sets M. ?P f A = ?P g A" + by (simp add: emeasure_density[symmetric]) 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" @@ -844,7 +845,7 @@ unfolding indicator_def by auto 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: *) + (auto intro!: borel_measurable_ereal_times f measure_eqI Int simp: emeasure_density * subset_eq) 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" @@ -933,7 +934,42 @@ 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: +lemma sigma_finite_density_unique: + assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" + assumes pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ g x" + and fin: "sigma_finite_measure (density M f)" + shows "density M f = density M g \ (AE x in M. f x = g x)" +proof + assume "AE x in M. f x = g x" with borel show "density M f = density M g" + by (auto intro: density_cong) +next + assume eq: "density M f = density M g" + interpret f!: sigma_finite_measure "density M f" by fact + from f.sigma_finite_incseq guess A . note cover = this + + have "AE x in M. \i. x \ A i \ f x = g x" + unfolding AE_all_countable + proof + fix i + have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))" + unfolding eq .. + moreover have "(\\<^isup>+x. f x * indicator (A i) x \M) \ \" + using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq) + ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x" + using borel pos cover(1) pos + by (intro finite_density_unique[THEN iffD1]) + (auto simp: density_density_eq subset_eq) + then show "AE x in M. x \ A i \ f x = g x" + by auto + qed + with AE_space show "AE x in M. f x = g x" + apply eventually_elim + using cover(2)[symmetric] + apply auto + done +qed + +lemma (in sigma_finite_measure) sigma_finite_iff_density_finite': 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 \ _") @@ -1019,6 +1055,13 @@ qed qed +lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: + "f \ borel_measurable M \ sigma_finite_measure (density M f) \ (AE x in M. f x \ \)" + apply (subst density_max_0) + apply (subst sigma_finite_iff_density_finite') + apply (auto simp: max_def intro!: measurable_If) + done + section "Radon-Nikodym derivative" definition @@ -1069,15 +1112,17 @@ 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+ + unfolding eq[symmetric] + by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv_density + RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) + +lemma RN_deriv_unique_sigma_finite: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and eq: "density M f = N" and fin: "sigma_finite_measure N" + shows "AE x in M. f x = RN_deriv M N x" + using fin unfolding eq[symmetric] + by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv_density + RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) lemma (in sigma_finite_measure) RN_deriv_distr: fixes T :: "'a \ 'b" diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Wed Oct 10 13:30:50 2012 +0200 @@ -1292,11 +1292,11 @@ lemma measurable_If_set: assumes measure: "f \ measurable M M'" "g \ measurable M M'" - assumes P: "A \ sets M" + assumes P: "A \ space M \ 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 + have "{x \ space M. x \ A} = A \ space M" by auto + thus "{x \ space M. x \ A} \ sets M" using `A \ space M \ sets M` by auto qed lemma measurable_ident[intro, simp]: "id \ measurable M M" @@ -1310,6 +1310,10 @@ apply force+ done +lemma measurable_compose: + "f \ measurable M N \ g \ measurable N L \ (\x. g (f x)) \ measurable M L" + using measurable_comp[of f M N g L] by (simp add: comp_def) + 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" @@ -2037,4 +2041,25 @@ using assms by (auto simp: dynkin_def) qed +lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]: + assumes "Int_stable G" + and closed: "G \ Pow \" + and A: "A \ sigma_sets \ G" + assumes basic: "\A. A \ G \ P A" + and empty: "P {}" + and compl: "\A. A \ sigma_sets \ G \ P A \ P (\ - A)" + and union: "\A. disjoint_family A \ range A \ sigma_sets \ G \ (\i. P (A i)) \ P (\i::nat. A i)" + shows "P A" +proof - + let ?D = "{ A \ sigma_sets \ G. P A }" + interpret sigma_algebra \ "sigma_sets \ G" + using closed by (rule sigma_algebra_sigma_sets) + from compl[OF _ empty] closed have space: "P \" by simp + interpret dynkin_system \ ?D + by default (auto dest: sets_into_space intro!: space compl union) + have "sigma_sets \ G = ?D" + by (rule dynkin_lemma) (auto simp: basic `Int_stable G`) + with A show ?thesis by auto +qed + end diff -r b1493798d253 -r 9a2a53be24a2 src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Oct 10 13:04:15 2012 +0200 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Wed Oct 10 13:30:50 2012 +0200 @@ -145,7 +145,7 @@ "msgs = keys \ {ms. set ms \ messages \ length ms = n}" definition P :: "('key \ 'message list) \ real" where - "P = (\(k, ms). K k * (\i(k, ms). K k * (\i(k, ms). map (observe k) ms)" definition t :: "'observation list \ 'observation \ nat" where - "t seq obs = length (filter (op = obs) seq)" + t_def2: "t seq obs = card { i. i < length seq \ seq ! i = obs}" + +lemma t_def: "t seq obs = length (filter (op = obs) seq)" + unfolding t_def2 length_filter_conv_card by (subst eq_commute) simp lemma card_T_bound: "card ((t\OB)`msgs) \ (n+1)^card observations" proof - @@ -306,7 +309,6 @@ 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 - @@ -325,7 +327,7 @@ -(\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)" +lemma ce_OB_eq_ce_t: "\(fst ; OB) = \(fst ; t\OB)" proof - txt {* Lemma 2 *} @@ -434,7 +436,6 @@ 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 using * by simp also have "\ = -(\obs\t`OB`msgs. \

(t\OB) {obs} * ?Ht obs)" @@ -452,16 +453,15 @@ also have "\ = \(fst | t\OB)" unfolding conditional_entropy_eq_ce_with_hypothesis by (simp add: comp_def image_image[symmetric]) - finally show ?thesis . + finally show ?thesis + by (subst (1 2) mutual_information_eq_entropy_conditional_entropy) simp_all qed theorem "\(fst ; OB) \ real (card observations) * log b (real n + 1)" proof - - from simple_function_finite simple_function_finite - have "\(fst ; OB) = \(fst) - \(fst | OB)" - by (rule mutual_information_eq_entropy_conditional_entropy) - also have "\ = \(fst) - \(fst | t\OB)" - unfolding ce_OB_eq_ce_t .. + have "\(fst ; OB) = \(fst) - \(fst | t\OB)" + unfolding ce_OB_eq_ce_t + by (rule mutual_information_eq_entropy_conditional_entropy simple_function_finite)+ 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) simp