diff -r 73e2d802ea41 -r d8f3fc934ff6 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Thu May 26 14:12:06 2011 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Thu May 26 17:40:01 2011 +0200 @@ -325,6 +325,10 @@ "\ \i. i \ I \ A i \ sets (M i) \ \ Pi\<^isub>E I A \ sets P" by (auto simp: sets_product_algebra) +lemma Int_stable_product_algebra_generator: + "(\i. i \ I \ Int_stable (M i)) \ Int_stable (product_algebra_generator I M)" + by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff) + section "Generating set generates also product algebra" lemma sigma_product_algebra_sigma_eq: @@ -478,6 +482,32 @@ using `A \ sets (M i)` by (auto intro!: product_algebraI) qed (insert `i \ I`, auto) +lemma (in sigma_algebra) measurable_restrict: + assumes I: "finite I" + assumes "\i. i \ I \ sets (N i) \ Pow (space (N i))" + assumes X: "\i. i \ I \ X i \ measurable M (N i)" + shows "(\x. \i\I. X i x) \ measurable M (Pi\<^isub>M I N)" + unfolding product_algebra_def +proof (simp, rule measurable_sigma) + show "sets (product_algebra_generator I N) \ Pow (space (product_algebra_generator I N))" + by (rule product_algebra_generator_sets_into_space) fact + show "(\x. \i\I. X i x) \ space M \ space (product_algebra_generator I N)" + using X by (auto simp: measurable_def) + fix E assume "E \ sets (product_algebra_generator I N)" + then obtain F where "E = Pi\<^isub>E I F" and F: "\i. i \ I \ F i \ sets (N i)" + by (auto simp: product_algebra_generator_def) + then have "(\x. \i\I. X i x) -` E \ space M = (\i\I. X i -` F i \ space M) \ space M" + by (auto simp: Pi_iff) + also have "\ \ sets M" + proof cases + assume "I = {}" then show ?thesis by simp + next + assume "I \ {}" with X F I show ?thesis + by (intro finite_INT measurable_sets Int) auto + qed + finally show "(\x. \i\I. X i x) -` E \ space M \ sets M" . +qed + locale product_sigma_finite = fixes M :: "'i \ ('a,'b) measure_space_scheme" assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" @@ -719,9 +749,8 @@ using A unfolding product_algebra_def by auto next show "Int_stable IJ.G" - by (simp add: PiE_Int Int_stable_def product_algebra_def - product_algebra_generator_def) - auto + by (rule Int_stable_product_algebra_generator) + (auto simp: Int_stable_def) show "range ?F \ sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def product_algebra_generator_def)