diff -r c4b809e57fe0 -r 96984bf6fa5b src/HOL/Probability/Borel.thy --- a/src/HOL/Probability/Borel.thy Fri Aug 27 13:48:10 2010 +0200 +++ b/src/HOL/Probability/Borel.thy Fri Aug 27 14:05:03 2010 +0200 @@ -658,6 +658,30 @@ "(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_measureable_euclidean_component: + "(\x::'a::euclidean_space. x $$ i) \ borel_measurable borel_space" + unfolding borel_space_def[where 'a=real] +proof (rule borel_space.measurable_sigma) + fix S::"real set" assume "S \ open" then have "open S" unfolding mem_def . + from open_vimage_euclidean_component[OF this] + show "(\x. x $$ i) -` S \ space borel_space \ sets borel_space" + by (auto intro: borel_space_open) +qed auto + +lemma (in sigma_algebra) borel_measureable_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_measureable_euclidean_component) +next + assume f: "\ix. f x $$ i) \ borel_measurable M" + then show "f \ borel_measurable M" + unfolding borel_measurable_iff_halfspace_le by auto +qed + subsection "Borel measurable operators" lemma (in sigma_algebra) affine_borel_measurable_vector: