# HG changeset patch # User hoelzl # Date 1301401661 -7200 # Node ID 61d5d50ca74c2cfc50629eceb0b745bb6e310e29 # Parent 5b52c6a9c62741cf0a5cd8bdb72ff29b6fdf0314 add infinite product measure diff -r 5b52c6a9c627 -r 61d5d50ca74c src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Mar 29 14:27:39 2011 +0200 +++ b/src/HOL/IsaMakefile Tue Mar 29 14:27:41 2011 +0200 @@ -1190,7 +1190,8 @@ Probability/Caratheodory.thy Probability/Complete_Measure.thy \ Probability/ex/Dining_Cryptographers.thy \ Probability/ex/Koepf_Duermuth_Countermeasure.thy \ - Probability/Finite_Product_Measure.thy Probability/Information.thy \ + Probability/Finite_Product_Measure.thy \ + Probability/Infinite_Product_Measure.thy Probability/Information.thy \ Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \ Probability/Measure.thy Probability/Probability_Space.thy \ Probability/Probability.thy Probability/Radon_Nikodym.thy \ diff -r 5b52c6a9c627 -r 61d5d50ca74c src/HOL/Probability/Infinite_Product_Measure.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Tue Mar 29 14:27:41 2011 +0200 @@ -0,0 +1,753 @@ +(* Title: HOL/Probability/Infinite_Product_Measure.thy + Author: Johannes Hölzl, TU München +*) + +header {*Infinite Product Measure*} + +theory Infinite_Product_Measure + imports Probability_Space +begin + +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)" + unfolding merge_def by auto + +lemma extensional_merge_sub: "I \ J \ K \ merge I x J y \ extensional K" + unfolding merge_def extensional_def by auto + +lemma injective_vimage_restrict: + assumes J: "J \ I" + and sets: "A \ (\\<^isub>E i\J. S i)" "B \ (\\<^isub>E i\J. S i)" and ne: "(\\<^isub>E i\I. S i) \ {}" + and eq: "(\x. restrict x J) -` A \ (\\<^isub>E i\I. S i) = (\x. restrict x J) -` B \ (\\<^isub>E i\I. S i)" + shows "A = B" +proof (intro set_eqI) + fix x + from ne obtain y where y: "\i. i \ I \ y i \ S i" by auto + have "J \ (I - J) = {}" by auto + 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)" + 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) + next + assume "x \ (\\<^isub>E i\J. S i)" with sets show "x \ A \ x \ B" by auto + qed +qed + +locale product_prob_space = + fixes M :: "'i \ ('a,'b) measure_space_scheme" and I :: "'i set" + assumes prob_spaces: "\i. prob_space (M i)" + and I_not_empty: "I \ {}" + +locale finite_product_prob_space = product_prob_space M I + for M :: "'i \ ('a,'b) measure_space_scheme" and I :: "'i set" + + assumes finite_index'[intro]: "finite I" + +sublocale product_prob_space \ M: prob_space "M i" for i + by (rule prob_spaces) + +sublocale product_prob_space \ product_sigma_finite + by default + +sublocale finite_product_prob_space \ finite_product_sigma_finite + by default (fact finite_index') + +sublocale finite_product_prob_space \ prob_space "Pi\<^isub>M I M" +proof + show "measure P (space P) = 1" + by (simp add: measure_times measure_space_1 setprod_1) +qed + +lemma (in product_prob_space) measure_preserving_restrict: + assumes "J \ {}" "J \ K" "finite K" + shows "(\f. restrict f J) \ measure_preserving (\\<^isub>M i\K. M i) (\\<^isub>M i\J. M i)" (is "?R \ _") +proof - + interpret K: finite_product_prob_space M K + by default (insert assms, auto) + have J: "J \ {}" "finite J" using assms by (auto simp add: finite_subset) + interpret J: finite_product_prob_space M J + by default (insert J, auto) + from J.sigma_finite_pairs guess F .. note F = this + then have [simp,intro]: "\k i. k \ J \ F k i \ sets (M k)" + by auto + let "?F i" = "\\<^isub>E k\J. F k i" + let ?J = "product_algebra_generator J M \ measure := measure (Pi\<^isub>M J M) \" + have "?R \ measure_preserving (\\<^isub>M i\K. M i) (sigma ?J)" + proof (rule K.measure_preserving_Int_stable) + show "Int_stable ?J" + by (auto simp: Int_stable_def product_algebra_generator_def PiE_Int) + show "range ?F \ sets ?J" "incseq ?F" "(\i. ?F i) = space ?J" + using F by auto + show "\i. measure ?J (?F i) \ \" + using F by (simp add: J.measure_times setprod_PInf) + have "measure_space (Pi\<^isub>M J M)" by default + then show "measure_space (sigma ?J)" + by (simp add: product_algebra_def sigma_def) + show "?R \ measure_preserving (Pi\<^isub>M K M) ?J" + proof (simp add: measure_preserving_def measurable_def product_algebra_generator_def del: vimage_Int, + safe intro!: restrict_extensional) + fix x k assume "k \ J" "x \ (\ i\K. space (M i))" + then show "x k \ space (M k)" using `J \ K` by auto + next + fix E assume "E \ (\ i\J. sets (M i))" + then have E: "\j. j \ J \ E j \ sets (M j)" by auto + then have *: "?R -` Pi\<^isub>E J E \ (\\<^isub>E i\K. space (M i)) = (\\<^isub>E i\K. if i \ J then E i else space (M i))" + (is "?X = Pi\<^isub>E K ?M") + using `J \ K` sets_into_space by (auto simp: Pi_iff split: split_if_asm) blast+ + with E show "?X \ sets (Pi\<^isub>M K M)" + by (auto intro!: product_algebra_generatorI) + have "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = (\i\J. measure (M i) (?M i))" + using E by (simp add: J.measure_times) + also have "\ = measure (Pi\<^isub>M K M) ?X" + unfolding * using E `finite K` `J \ K` + by (auto simp: K.measure_times M.measure_space_1 + cong del: setprod_cong + intro!: setprod_mono_one_left) + finally show "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = measure (Pi\<^isub>M K M) ?X" . + qed + qed + then show ?thesis + by (simp add: product_algebra_def sigma_def) +qed + +lemma (in product_prob_space) measurable_restrict: + assumes *: "J \ {}" "J \ K" "finite K" + shows "(\f. restrict f J) \ measurable (\\<^isub>M i\K. M i) (\\<^isub>M i\J. M i)" + using measure_preserving_restrict[OF *] + by (rule measure_preservingD2) + +definition (in product_prob_space) + "emb J K X = (\x. restrict x K) -` X \ space (Pi\<^isub>M J M)" + +lemma (in product_prob_space) emb_trans[simp]: + "J \ K \ K \ L \ emb L K (emb K J X) = emb L J X" + by (auto simp add: Int_absorb1 emb_def) + +lemma (in product_prob_space) emb_empty[simp]: + "emb K J {} = {}" + by (simp add: emb_def) + +lemma (in product_prob_space) emb_Pi: + assumes "X \ (\ j\J. sets (M j))" "J \ K" + shows "emb K J (Pi\<^isub>E J X) = (\\<^isub>E i\K. if i \ J then X i else space (M i))" + using assms space_closed + by (auto simp: emb_def Pi_iff split: split_if_asm) blast+ + +lemma (in product_prob_space) emb_injective: + assumes "J \ {}" "J \ L" "finite J" and sets: "X \ sets (Pi\<^isub>M J M)" "Y \ sets (Pi\<^isub>M J M)" + assumes "emb L J X = emb L J Y" + shows "X = Y" +proof - + interpret J: finite_product_sigma_finite M J by default fact + show "X = Y" + proof (rule injective_vimage_restrict) + show "X \ (\\<^isub>E i\J. space (M i))" "Y \ (\\<^isub>E i\J. space (M i))" + using J.sets_into_space sets by auto + have "\i\L. \x. x \ space (M i)" + using M.not_empty by auto + from bchoice[OF this] + show "(\\<^isub>E i\L. space (M i)) \ {}" by auto + show "(\x. restrict x J) -` X \ (\\<^isub>E i\L. space (M i)) = (\x. restrict x J) -` Y \ (\\<^isub>E i\L. space (M i))" + using `emb L J X = emb L J Y` by (simp add: emb_def) + qed fact +qed + +lemma (in product_prob_space) emb_id: + "B \ (\\<^isub>E i\L. space (M i)) \ emb L L B = B" + by (auto simp: emb_def Pi_iff subset_eq extensional_restrict) + +lemma (in product_prob_space) emb_simps: + shows "emb L K (A \ B) = emb L K A \ emb L K B" + and "emb L K (A \ B) = emb L K A \ emb L K B" + and "emb L K (A - B) = emb L K A - emb L K B" + by (auto simp: emb_def) + +lemma (in product_prob_space) measurable_emb[intro,simp]: + assumes *: "J \ {}" "J \ L" "finite L" "X \ sets (Pi\<^isub>M J M)" + shows "emb L J X \ sets (Pi\<^isub>M L M)" + using measurable_restrict[THEN measurable_sets, OF *] by (simp add: emb_def) + +lemma (in product_prob_space) measure_emb[intro,simp]: + assumes *: "J \ {}" "J \ L" "finite L" "X \ sets (Pi\<^isub>M J M)" + shows "measure (Pi\<^isub>M L M) (emb L J X) = measure (Pi\<^isub>M J M) X" + using measure_preserving_restrict[THEN measure_preservingD, OF *] + by (simp add: emb_def) + +definition (in product_prob_space) generator :: "('i \ 'a) measure_space" where + "generator = \ + space = (\\<^isub>E i\I. space (M i)), + sets = (\J\{J. J \ {} \ finite J \ J \ I}. emb I J ` sets (Pi\<^isub>M J M)), + measure = undefined + \" + +lemma (in product_prob_space) generatorI: + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ A \ sets generator" + unfolding generator_def by auto + +lemma (in product_prob_space) generatorI': + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ emb I J X \ sets generator" + unfolding generator_def by auto + +lemma (in product_sigma_finite) + assumes "I \ J = {}" "finite I" "finite J" and A: "A \ sets (Pi\<^isub>M (I \ J) M)" + shows measure_fold_integral: + "measure (Pi\<^isub>M (I \ J) M) A = (\\<^isup>+x. measure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M)) \Pi\<^isub>M I M)" (is ?I) + and measure_fold_measurable: + "(\x. measure (Pi\<^isub>M J M) (merge I x J -` A \ space (Pi\<^isub>M J M))) \ borel_measurable (Pi\<^isub>M I M)" (is ?B) +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + interpret IJ: pair_sigma_finite I.P J.P .. + show ?I + unfolding measure_fold[OF assms] + apply (subst IJ.pair_measure_alt) + apply (intro measurable_sets[OF _ A] measurable_merge assms) + apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure + intro!: I.positive_integral_cong) + done + + have "(\(x, y). merge I x J y) -` A \ space (I.P \\<^isub>M J.P) \ sets (I.P \\<^isub>M J.P)" + by (intro measurable_sets[OF _ A] measurable_merge assms) + from IJ.measure_cut_measurable_fst[OF this] + show ?B + apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure) + apply (subst (asm) measurable_cong) + apply auto + done +qed + +lemma (in prob_space) measure_le_1: "X \ sets M \ \ X \ 1" + unfolding measure_space_1[symmetric] + using sets_into_space + by (intro measure_mono) auto + +definition (in product_prob_space) + "\G A = + (THE x. \J. J \ {} \ finite J \ J \ I \ (\X\sets (Pi\<^isub>M J M). A = emb I J X \ x = measure (Pi\<^isub>M J M) X))" + +lemma (in product_prob_space) \G_spec: + assumes J: "J \ {}" "finite J" "J \ I" "A = emb I J X" "X \ sets (Pi\<^isub>M J M)" + shows "\G A = measure (Pi\<^isub>M J M) X" + unfolding \G_def +proof (intro the_equality allI impI ballI) + fix K Y assume K: "K \ {}" "finite K" "K \ I" "A = emb I K Y" "Y \ sets (Pi\<^isub>M K M)" + have "measure (Pi\<^isub>M K M) Y = measure (Pi\<^isub>M (K \ J) M) (emb (K \ J) K Y)" + using K J by simp + also have "emb (K \ J) K Y = emb (K \ J) J X" + using K J by (simp add: emb_injective[of "K \ J" I]) + also have "measure (Pi\<^isub>M (K \ J) M) (emb (K \ J) J X) = measure (Pi\<^isub>M J M) X" + using K J by simp + finally show "measure (Pi\<^isub>M J M) X = measure (Pi\<^isub>M K M) Y" .. +qed (insert J, force) + +lemma (in product_prob_space) \G_eq: + "J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ \G (emb I J X) = measure (Pi\<^isub>M J M) X" + by (intro \G_spec) auto + +lemma (in product_prob_space) generator_Ex: + assumes *: "A \ sets generator" + shows "\J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A = emb I J X \ \G A = measure (Pi\<^isub>M J M) X" +proof - + from * obtain J X where J: "J \ {}" "finite J" "J \ I" "A = emb I J X" "X \ sets (Pi\<^isub>M J M)" + unfolding generator_def by auto + with \G_spec[OF this] show ?thesis by auto +qed + +lemma (in product_prob_space) generatorE: + assumes A: "A \ sets generator" + obtains J X where "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" "emb I J X = A" "\G A = measure (Pi\<^isub>M J M) X" +proof - + from generator_Ex[OF A] obtain X J where "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" "emb I J X = A" + "\G A = measure (Pi\<^isub>M J M) X" by auto + then show thesis by (intro that) auto +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 - + interpret J: finite_product_sigma_algebra M J by default fact + interpret K: finite_product_sigma_algebra M K by default fact + interpret JK: pair_sigma_algebra J.P K.P .. + + from JK.measurable_cut_fst[OF + 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 + +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))" +proof - + have [simp]: "\x J K L. merge J y K (restrict x L) = merge J y (K \ L) 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" + 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 + have [simp]: "(K - J) \ K = K - J" by auto + from y `K \ I` `J \ I` show ?thesis + by (simp split: split_merge add: emb_def Pi_iff extensional_merge_sub set_eq_iff) auto +qed + +definition (in product_prob_space) infprod_algebra :: "('i \ 'a) measure_space" where + "infprod_algebra = sigma generator \ measure := + (SOME \. (\s\sets generator. \ s = \G s) \ + measure_space \space = space generator, sets = sets (sigma generator), measure = \\)\" + +syntax + "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3PIP _:_./ _)" 10) + +syntax (xsymbols) + "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3\\<^isub>P _\_./ _)" 10) + +syntax (HTML output) + "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme" ("(3\\<^isub>P _\_./ _)" 10) + +abbreviation + "Pi\<^isub>P I M \ product_prob_space.infprod_algebra M I" + +translations + "PIP x:I. M" == "CONST Pi\<^isub>P I (%x. M)" + +sublocale product_prob_space \ G: algebra generator +proof + let ?G = generator + show "sets ?G \ Pow (space ?G)" + by (auto simp: generator_def emb_def) + from I_not_empty + obtain i where "i \ I" by auto + then show "{} \ sets ?G" + by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\i. {}"] + simp: product_algebra_def sigma_def sigma_sets.Empty generator_def emb_def) + from `i \ I` show "space ?G \ sets ?G" + by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\i. space (M i))"] + simp: generator_def emb_def) + fix A assume "A \ sets ?G" + then obtain JA XA where XA: "JA \ {}" "finite JA" "JA \ I" "XA \ sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA" + by (auto simp: generator_def) + fix B assume "B \ sets ?G" + then obtain JB XB where XB: "JB \ {}" "finite JB" "JB \ I" "XB \ sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB" + by (auto simp: generator_def) + let ?RA = "emb (JA \ JB) JA XA" + let ?RB = "emb (JA \ JB) JB XB" + interpret JAB: finite_product_sigma_algebra M "JA \ JB" + by default (insert XA XB, auto) + have *: "A - B = emb I (JA \ JB) (?RA - ?RB)" "A \ B = emb I (JA \ JB) (?RA \ ?RB)" + using XA A XB B by (auto simp: emb_simps) + then show "A - B \ sets ?G" "A \ B \ sets ?G" + using XA XB by (auto intro!: generatorI') +qed + +lemma (in product_prob_space) positive_\G: "positive generator \G" +proof (intro positive_def[THEN iffD2] conjI ballI) + from generatorE[OF G.empty_sets] guess J X . note this[simp] + interpret J: finite_product_sigma_finite M J by default fact + have "X = {}" + by (rule emb_injective[of J I]) simp_all + then show "\G {} = 0" by simp +next + fix A assume "A \ sets generator" + from generatorE[OF this] guess J X . note this[simp] + interpret J: finite_product_sigma_finite M J by default fact + show "0 \ \G A" by simp +qed + +lemma (in product_prob_space) additive_\G: "additive generator \G" +proof (intro additive_def[THEN iffD2] ballI impI) + fix A assume "A \ sets generator" with generatorE guess J X . note J = this + fix B assume "B \ sets generator" with generatorE guess K Y . note K = this + assume "A \ B = {}" + have JK: "J \ K \ {}" "J \ K \ I" "finite (J \ K)" + using J K by auto + interpret JK: finite_product_sigma_finite M "J \ K" by default fact + have JK_disj: "emb (J \ K) J X \ emb (J \ K) K Y = {}" + apply (rule emb_injective[of "J \ K" I]) + apply (insert `A \ B = {}` JK J K) + apply (simp_all add: JK.Int emb_simps) + done + have AB: "A = emb I (J \ K) (emb (J \ K) J X)" "B = emb I (J \ K) (emb (J \ K) K Y)" + using J K by simp_all + then have "\G (A \ B) = \G (emb I (J \ K) (emb (J \ K) J X \ emb (J \ K) K Y))" + by (simp add: emb_simps) + also have "\ = measure (Pi\<^isub>M (J \ K) M) (emb (J \ K) J X \ emb (J \ K) K Y)" + using JK J(1, 4) K(1, 4) by (simp add: \G_eq JK.Un) + also have "\ = \G A + \G B" + using J K JK_disj by (simp add: JK.measure_additive[symmetric]) + finally show "\G (A \ B) = \G A + \G B" . +qed + +lemma (in product_prob_space) finite_index_eq_finite_product: + assumes "finite I" + shows "sets (sigma generator) = sets (Pi\<^isub>M I M)" +proof safe + interpret I: finite_product_sigma_algebra M I by default fact + have [simp]: "space generator = space (Pi\<^isub>M I M)" + by (simp add: generator_def product_algebra_def) + { fix A assume "A \ sets (sigma generator)" + then show "A \ sets I.P" unfolding sets_sigma + proof induct + case (Basic A) + from generatorE[OF this] guess J X . note J = this + with `finite I` have "emb I J X \ sets I.P" by auto + with `emb I J X = A` show "A \ sets I.P" by simp + qed auto } + { fix A assume "A \ sets I.P" + moreover with I.sets_into_space have "emb I I A = A" by (intro emb_id) auto + ultimately show "A \ sets (sigma generator)" + using `finite I` I_not_empty unfolding sets_sigma + by (intro sigma_sets.Basic generatorI[of I A]) auto } +qed + +lemma (in product_prob_space) extend_\G: + "\\. (\s\sets generator. \ s = \G s) \ + measure_space \space = space generator, sets = sets (sigma generator), measure = \\" +proof cases + assume "finite I" + interpret I: finite_product_sigma_finite M I by default fact + show ?thesis + proof (intro exI[of _ "measure (Pi\<^isub>M I M)"] ballI conjI) + fix A assume "A \ sets generator" + from generatorE[OF this] guess J X . note J = this + from J(1-4) `finite I` show "measure I.P A = \G A" + unfolding J(6) + by (subst J(5)[symmetric]) (simp add: measure_emb) + next + have [simp]: "space generator = space (Pi\<^isub>M I M)" + by (simp add: generator_def product_algebra_def) + have "\space = space generator, sets = sets (sigma generator), measure = measure I.P\ + = I.P" (is "?P = _") + by (auto intro!: measure_space.equality simp: finite_index_eq_finite_product[OF `finite I`]) + then show "measure_space ?P" by simp default + qed +next + let ?G = generator + assume "\ finite I" + note \G_mono = + G.additive_increasing[OF positive_\G additive_\G, THEN increasingD] + + { fix Z J assume J: "J \ {}" "finite J" "J \ I" and Z: "Z \ sets ?G" + + from `infinite I` `finite J` obtain k where k: "k \ I" "k \ J" + by (metis rev_finite_subset subsetI) + moreover from Z guess K' X' by (rule generatorE) + moreover def K \ "insert k K'" + moreover def X \ "emb K K' X'" + ultimately have K: "K \ {}" "finite K" "K \ I" "X \ sets (Pi\<^isub>M K M)" "Z = emb I K X" + "K - J \ {}" "K - J \ I" "\G Z = measure (Pi\<^isub>M K M) X" + 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)" + { 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)) \ sets ?G" + using J K by (intro generatorI) auto + have "\G (merge J y (I - J) -` emb I K X \ space (Pi\<^isub>M I M)) = measure (Pi\<^isub>M (K - J) M) (?M y)" + unfolding * using K J by (subst \G_eq[OF _ _ _ **]) auto + note * ** *** this } + note merge_in_G = this + + have "finite (K - J)" using K by auto + + interpret J: finite_product_prob_space M J by default fact+ + interpret KmJ: finite_product_prob_space M "K - J" by default fact+ + + have "\G Z = measure (Pi\<^isub>M (J \ (K - J)) M) (emb (J \ (K - J)) K X)" + using K J by simp + also have "\ = (\\<^isup>+ x. measure (Pi\<^isub>M (K - J) M) (?M x) \Pi\<^isub>M J M)" + using K J by (subst measure_fold_integral) auto + also have "\ = (\\<^isup>+ y. \G (merge J y (I - J) -` Z \ space (Pi\<^isub>M I M)) \Pi\<^isub>M J M)" + (is "_ = (\\<^isup>+x. \G (?MZ x) \Pi\<^isub>M J M)") + proof (intro J.positive_integral_cong) + fix x assume x: "x \ space (Pi\<^isub>M J M)" + with K merge_in_G(2)[OF this] + show "measure (Pi\<^isub>M (K - J) M) (?M x) = \G (?MZ x)" + unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \G_eq) auto + qed + finally have fold: "\G Z = (\\<^isup>+x. \G (?MZ x) \Pi\<^isub>M J M)" . + + { fix x assume x: "x \ space (Pi\<^isub>M J M)" + then have "\G (?MZ x) \ 1" + unfolding merge_in_G(4)[OF x] `Z = emb I K X` + 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))" + have "?q \ borel_measurable (Pi\<^isub>M J M)" + unfolding `Z = emb I K X` using J K merge_in_G(3) + by (simp add: merge_in_G \G_eq measure_fold_measurable + del: space_product_algebra cong: measurable_cong) + note this fold le_1 merge_in_G(3) } + note fold = this + + show ?thesis + proof (rule G.caratheodory_empty_continuous[OF positive_\G additive_\G]) + fix A assume "A \ sets ?G" + with generatorE guess J X . note JX = this + interpret JK: finite_product_prob_space M J by default fact+ + with JX show "\G A \ \" by simp + next + fix A assume A: "range A \ sets ?G" "decseq A" "(\i. A i) = {}" + then have "decseq (\i. \G (A i))" + by (auto intro!: \G_mono simp: decseq_def) + moreover + have "(INF i. \G (A i)) = 0" + proof (rule ccontr) + assume "(INF i. \G (A i)) \ 0" (is "?a \ 0") + moreover have "0 \ ?a" + using A positive_\G by (auto intro!: le_INFI simp: positive_def) + ultimately have "0 < ?a" by auto + + have "\n. \J X. J \ {} \ finite J \ J \ I \ X \ sets (Pi\<^isub>M J M) \ A n = emb I J X \ \G (A n) = measure (Pi\<^isub>M J M) X" + using A by (intro allI generator_Ex) auto + then obtain J' X' where J': "\n. J' n \ {}" "\n. finite (J' n)" "\n. J' n \ I" "\n. X' n \ sets (Pi\<^isub>M (J' n) M)" + and A': "\n. A n = emb I (J' n) (X' n)" + unfolding choice_iff by blast + moreover def J \ "\n. (\i\n. J' i)" + moreover def X \ "\n. emb (J n) (J' n) (X' n)" + ultimately have J: "\n. J n \ {}" "\n. finite (J n)" "\n. J n \ I" "\n. X n \ sets (Pi\<^isub>M (J n) M)" + by auto + with A' have A_eq: "\n. A n = emb I (J n) (X n)" "\n. A n \ sets ?G" + unfolding J_def X_def by (subst emb_trans) (insert A, auto) + + have J_mono: "\n m. n \ m \ J n \ J m" + unfolding J_def by force + + interpret J: finite_product_prob_space M "J i" for i by default fact+ + + have a_le_1: "?a \ 1" + using \G_spec[of "J 0" "A 0" "X 0"] J A_eq + by (auto intro!: INF_leI2[of 0] J.measure_le_1) + + let "?M K Z y" = "merge K y (I - K) -` Z \ space (Pi\<^isub>M I M)" + + { fix Z k assume Z: "range Z \ sets ?G" "decseq Z" "\n. ?a / 2^k \ \G (Z n)" + then have Z_sets: "\n. Z n \ sets ?G" by auto + fix J' assume J': "J' \ {}" "finite J'" "J' \ I" + interpret J': finite_product_prob_space M J' by default fact+ + + let "?q n y" = "\G (?M J' (Z n) y)" + let "?Q n" = "?q n -` {?a / 2^(k+1) ..} \ space (Pi\<^isub>M J' M)" + { fix n + have "?q n \ borel_measurable (Pi\<^isub>M J' M)" + using Z J' by (intro fold(1)) auto + then have "?Q n \ sets (Pi\<^isub>M J' M)" + by (rule measurable_sets) auto } + note Q_sets = this + + have "?a / 2^(k+1) \ (INF n. measure (Pi\<^isub>M J' M) (?Q n))" + proof (intro le_INFI) + fix n + have "?a / 2^k \ \G (Z n)" using Z by auto + also have "\ \ (\\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \Pi\<^isub>M J' M)" + unfolding fold(2)[OF J' `Z n \ sets ?G`] + proof (intro J'.positive_integral_mono) + fix x assume x: "x \ space (Pi\<^isub>M J' M)" + then have "?q n x \ 1 + 0" + using J' Z fold(3) Z_sets by auto + also have "\ \ 1 + ?a / 2^(k+1)" + using `0 < ?a` by (intro add_mono) auto + finally have "?q n x \ 1 + ?a / 2^(k+1)" . + with x show "?q n x \ indicator (?Q n) x + ?a / 2^(k+1)" + by (auto split: split_indicator simp del: power_Suc) + qed + also have "\ = measure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)" + using `0 \ ?a` Q_sets J'.measure_space_1 + by (subst J'.positive_integral_add) auto + finally show "?a / 2^(k+1) \ measure (Pi\<^isub>M J' M) (?Q n)" using `?a \ 1` + by (cases rule: extreal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"]) + (auto simp: field_simps) + qed + also have "\ = measure (Pi\<^isub>M J' M) (\n. ?Q n)" + proof (intro J'.continuity_from_above) + show "range ?Q \ sets (Pi\<^isub>M J' M)" using Q_sets by auto + show "decseq ?Q" + unfolding decseq_def + proof (safe intro!: vimageI[OF refl]) + fix m n :: nat assume "m \ n" + fix x assume x: "x \ space (Pi\<^isub>M J' M)" + assume "?a / 2^(k+1) \ ?q n x" + also have "?q n x \ ?q m x" + proof (rule \G_mono) + from fold(4)[OF J', OF Z_sets x] + show "?M J' (Z n) x \ sets ?G" "?M J' (Z m) x \ sets ?G" by auto + show "?M J' (Z n) x \ ?M J' (Z m) x" + using `decseq Z`[THEN decseqD, OF `m \ n`] by auto + qed + finally show "?a / 2^(k+1) \ ?q m x" . + qed + qed (intro J'.finite_measure Q_sets) + finally have "(\n. ?Q n) \ {}" + using `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq) + then have "\w\space (Pi\<^isub>M J' M). \n. ?a / 2 ^ (k + 1) \ ?q n w" by auto } + note Ex_w = this + + let "?q k n y" = "\G (?M (J k) (A n) y)" + + have "\n. ?a / 2 ^ 0 \ \G (A n)" by (auto intro: INF_leI) + from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this + + let "?P k wk w" = + "w \ space (Pi\<^isub>M (J (Suc k)) M) \ restrict w (J k) = wk \ (\n. ?a / 2 ^ (Suc k + 1) \ ?q (Suc k) n w)" + def w \ "nat_rec w0 (\k wk. Eps (?P k wk))" + + { fix k have w: "w k \ space (Pi\<^isub>M (J k) M) \ + (\n. ?a / 2 ^ (k + 1) \ ?q k n (w k)) \ (k \ 0 \ restrict (w k) (J (k - 1)) = w (k - 1))" + proof (induct k) + case 0 with w0 show ?case + unfolding w_def nat_rec_0 by auto + next + case (Suc k) + then have wk: "w k \ space (Pi\<^isub>M (J k) M)" by auto + have "\w'. ?P k (w k) w'" + proof cases + assume [simp]: "J k = J (Suc k)" + show ?thesis + proof (intro exI[of _ "w k"] conjI allI) + fix n + have "?a / 2 ^ (Suc k + 1) \ ?a / 2 ^ (k + 1)" + using `0 < ?a` `?a \ 1` by (cases ?a) (auto simp: field_simps) + also have "\ \ ?q k n (w k)" using Suc by auto + finally show "?a / 2 ^ (Suc k + 1) \ ?q (Suc k) n (w k)" by simp + next + show "w k \ space (Pi\<^isub>M (J (Suc k)) M)" + using Suc by simp + then show "restrict (w k) (J k) = w k" + by (simp add: extensional_restrict) + qed + next + assume "J k \ J (Suc k)" + with J_mono[of k "Suc k"] have "J (Suc k) - J k \ {}" (is "?D \ {}") by auto + have "range (\n. ?M (J k) (A n) (w k)) \ sets ?G" + "decseq (\n. ?M (J k) (A n) (w k))" + "\n. ?a / 2 ^ (k + 1) \ \G (?M (J k) (A n) (w k))" + using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc + by (auto simp: decseq_def) + 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" + 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" + using w'(1) J(3)[of "Suc k"] + by (auto split: split_merge intro!: extensional_merge_sub) force+ + show ?thesis + apply (rule exI[of _ ?w]) + using w' J_mono[of k "Suc k"] wk unfolding * + apply (auto split: split_merge intro!: extensional_merge_sub ext) + apply (force simp: extensional_def) + done + qed + then have "?P k (w k) (w (Suc k))" + unfolding w_def nat_rec_Suc unfolding w_def[symmetric] + by (rule someI_ex) + then show ?case by auto + qed + moreover + then have "w k \ space (Pi\<^isub>M (J k) M)" by auto + moreover + from w have "?a / 2 ^ (k + 1) \ ?q k k (w k)" by auto + then have "?M (J k) (A k) (w k) \ {}" + using positive_\G[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 "\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) + ultimately have "w k \ space (Pi\<^isub>M (J k) M)" + "\x\A k. restrict x (J k) = w k" + "k \ 0 \ restrict (w k) (J (k - 1)) = w (k - 1)" + by auto } + note w = this + + { fix k l i assume "k \ l" "i \ J k" + { fix l have "w k i = w (k + l) i" + proof (induct l) + case (Suc l) + from `i \ J k` J_mono[of k "k + l"] have "i \ J (k + l)" by auto + with w(3)[of "k + Suc l"] + have "w (k + l) i = w (k + Suc l) i" + by (auto simp: restrict_def fun_eq_iff split: split_if_asm) + with Suc show ?case by simp + qed simp } + from this[of "l - k"] `k \ l` have "w l i = w k i" by simp } + note w_mono = this + + def w' \ "\i. if i \ (\k. J k) then w (LEAST k. i \ J k) i else if i \ I then (SOME x. x \ space (M i)) else undefined" + { fix i k assume k: "i \ J k" + have "w k i = w (LEAST k. i \ J k) i" + by (intro w_mono Least_le k LeastI[of _ k]) + then have "w' i = w k i" + unfolding w'_def using k by auto } + note w'_eq = this + have w'_simps1: "\i. i \ I \ w' i = undefined" + using J by (auto simp: w'_def) + have w'_simps2: "\i. i \ (\k. J k) \ i \ I \ w' i \ space (M i)" + using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]]) + { fix i assume "i \ I" then have "w' i \ space (M i)" + using w(1) by (cases "i \ (\k. J k)") (force simp: w'_simps2 w'_eq)+ } + note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this + + have w': "w' \ space (Pi\<^isub>M I M)" + using w(1) by (auto simp add: Pi_iff extensional_def) + + { fix n + have "restrict w' (J n) = w n" using w(1) + by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def) + with w[of n] obtain x where "x \ A n" "restrict x (J n) = restrict w' (J n)" by auto + then have "w' \ A n" unfolding A_eq using w' by (auto simp: emb_def) } + then have "w' \ (\i. A i)" by auto + with `(\i. A i) = {}` show False by auto + qed + ultimately show "(\i. \G (A i)) ----> 0" + using LIMSEQ_extreal_INFI[of "\i. \G (A i)"] by simp + qed +qed + +lemma (in product_prob_space) infprod_spec: + shows "(\s\sets generator. measure (Pi\<^isub>P I M) s = \G s) \ measure_space (Pi\<^isub>P I M)" +proof - + let ?P = "\\. (\A\sets generator. \ A = \G A) \ + measure_space \space = space generator, sets = sets (sigma generator), measure = \\" + have **: "measure infprod_algebra = (SOME \. ?P \)" + unfolding infprod_algebra_def by simp + have *: "Pi\<^isub>P I M = \space = space generator, sets = sets (sigma generator), measure = measure (Pi\<^isub>P I M)\" + unfolding infprod_algebra_def by auto + show ?thesis + apply (subst (2) *) + apply (unfold **) + apply (rule someI_ex[where P="?P"]) + apply (rule extend_\G) + done +qed + +sublocale product_prob_space \ measure_space "Pi\<^isub>P I M" + using infprod_spec by auto + +lemma (in product_prob_space) measure_infprod_emb: + assumes "J \ {}" "finite J" "J \ I" "X \ sets (Pi\<^isub>M J M)" + shows "measure (Pi\<^isub>P I M) (emb I J X) = measure (Pi\<^isub>M J M) X" +proof - + have "emb I J X \ sets generator" + using assms by (rule generatorI') + with \G_eq[OF assms] infprod_spec show ?thesis by auto +qed + +end \ No newline at end of file diff -r 5b52c6a9c627 -r 61d5d50ca74c src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Tue Mar 29 14:27:39 2011 +0200 +++ b/src/HOL/Probability/Probability.thy Tue Mar 29 14:27:41 2011 +0200 @@ -2,6 +2,8 @@ imports Complete_Measure Lebesgue_Measure + Probability + Infinite_Product_Measure Information "ex/Dining_Cryptographers" "ex/Koepf_Duermuth_Countermeasure"