# HG changeset patch # User hoelzl # Date 1401468485 -7200 # Node ID 7b31461802914980d480ce19cee4b3f683823f12 # Parent f174712d0a843ca25c6ffc20443dea152443f963 generalizd measurability on restricted space; rule for integrability on compact sets diff -r f174712d0a84 -r 7b3146180291 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Fri May 30 15:56:30 2014 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Fri May 30 18:48:05 2014 +0200 @@ -124,53 +124,68 @@ assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" -proof - - { fix X :: "ereal set" assume "X \ sets borel" - have 1: "(\x. f x * indicator \ x) -` X \ space M = - (if 0 \ X then (f -` X \ (\ \ space M)) \ (space M - (\ \ space M)) else f -` X \ (\ \ space M))" - by (auto split: split_if_asm simp: indicator_def) - have *: "f -` X \ (\ \ space M) = - (f -` X \ (\ \ space M) \ (space M - \ \ space M)) - (space M - \ \ space M)" - by auto - have "f -` X \ (\ \ space M) \ (space M - \ \ space M) \ sets M - \ f -` X \ (\ \ space M) \ sets M" - by (subst *) auto - note this 1 } - note 1 = this(1) and 2 = this(2) - - from 2 show ?thesis - by (auto simp add: measurable_def space_restrict_space sets_restrict_space_iff intro: 1 \) -qed + by (subst measurable_restrict_space_iff) + (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_cong) lemma borel_measurable_restrict_space_iff: fixes f :: "'a \ 'b::real_normed_vector" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. indicator \ x *\<^sub>R f x) \ borel_measurable M" -proof - - { fix X :: "'b set" assume "X \ sets borel" - have 1: "(\x. indicator \ x *\<^sub>R f x) -` X \ space M = - (if 0 \ X then (f -` X \ (\ \ space M)) \ (space M - (\ \ space M)) else f -` X \ (\ \ space M))" - by (auto split: split_if_asm simp: indicator_def) - have *: "f -` X \ (\ \ space M) = - (f -` X \ (\ \ space M) \ (space M - \ \ space M)) - (space M - \ \ space M)" - by auto - have "f -` X \ (\ \ space M) \ (space M - \ \ space M) \ sets M - \ f -` X \ (\ \ space M) \ sets M" - by (subst *) auto - note this 1 } - note 1 = this(1) and 2 = this(2) + by (subst measurable_restrict_space_iff) + (auto simp: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] mult_ac cong del: if_cong) + +lemma cbox_borel[measurable]: "cbox a b \ sets borel" + by (auto intro: borel_closed) + +lemma borel_compact: "compact (A::'a::t2_space set) \ A \ sets borel" + by (auto intro: borel_closed dest!: compact_imp_closed) - from 2 show ?thesis - by (auto simp add: measurable_def space_restrict_space sets_restrict_space_iff intro: 1 \) +lemma borel_measurable_continuous_on_if: + assumes A: "A \ sets borel" and f: "continuous_on A f" and g: "continuous_on (- A) g" + shows "(\x. if x \ A then f x else g x) \ borel_measurable borel" +proof (rule borel_measurableI) + fix S :: "'b set" assume "open S" + have "(\x. if x \ A then f x else g x) -` S \ space borel = (f -` S \ A) \ (g -` S \ -A)" + by (auto split: split_if_asm) + moreover obtain A' where "f -` S \ A = A' \ A" "open A'" + using continuous_on_open_invariant[THEN iffD1, rule_format, OF f `open S`] by metis + moreover obtain B' where "g -` S \ -A = B' \ -A" "open B'" + using continuous_on_open_invariant[THEN iffD1, rule_format, OF g `open S`] by metis + ultimately show "(\x. if x \ A then f x else g x) -` S \ space borel \ sets borel" + using A by auto qed 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 + "continuous_on UNIV f \ f \ borel_measurable borel" + using borel_measurable_continuous_on_if[of UNIV f "\_. undefined"] + by (auto intro: continuous_on_const) + +lemma borel_measurable_continuous_on: + 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': + "continuous_on A f \ open A \ + (\x. if x \ A then f x else c) \ borel_measurable borel" + using borel_measurable_continuous_on_if[of A f "\_. c"] by (auto intro: continuous_on_const) + +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_continuous_on_indicator: + fixes f g :: "'a::topological_space \ 'b::real_normed_vector" + assumes A: "A \ sets borel" and f: "continuous_on A f" + shows "(\x. indicator A x *\<^sub>R f x) \ borel_measurable borel" + using borel_measurable_continuous_on_if[OF assms, of "\_. 0"] + by (simp add: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] continuous_on_const + cong del: if_cong) lemma borel_eq_countable_basis: fixes B::"'a::topological_space set set" @@ -216,37 +231,6 @@ qed auto qed -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 - 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_continuous_Pair: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" assumes [measurable]: "f \ borel_measurable M" @@ -852,11 +836,8 @@ 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) - show "open ({0<..}::real set)" by auto - qed + by (rule borel_measurable_continuous_on_open[OF _ _ f]) + (auto intro!: continuous_intros) 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 . diff -r f174712d0a84 -r 7b3146180291 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Fri May 30 15:56:30 2014 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Fri May 30 18:48:05 2014 +0200 @@ -428,6 +428,16 @@ by (simp add: indicator_def [abs_def] cbox_interval[symmetric]) qed +lemma + fixes a b ::"'a::ordered_euclidean_space" + shows lmeasure_cbox[simp]: "emeasure lebesgue (cbox a b) = ereal (content (cbox a b))" +proof - + have "(indicator (UNIV \ {a..b})::_\real) integrable_on UNIV" + unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def] cbox_interval[symmetric]) + from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV + by (simp add: indicator_def [abs_def] cbox_interval[symmetric]) +qed + lemma lmeasure_singleton[simp]: fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0" using lmeasure_atLeastAtMost[of a a] by simp @@ -1035,40 +1045,56 @@ subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *} +lemma emeasure_bounded_finite: + assumes "bounded A" shows "emeasure lborel A < \" +proof - + from bounded_subset_cbox[OF `bounded A`] obtain a b where "A \ cbox a b" + by auto + then have "emeasure lborel A \ emeasure lborel (cbox a b)" + by (intro emeasure_mono) auto + then show ?thesis + by auto +qed + +lemma emeasure_compact_finite: "compact A \ emeasure lborel A < \" + using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded) + +lemma borel_integrable_compact: + fixes f :: "'a::ordered_euclidean_space \ 'b::{banach, second_countable_topology}" + assumes "compact S" "continuous_on S f" + shows "integrable lborel (\x. indicator S x *\<^sub>R f x)" +proof cases + assume "S \ {}" + have "continuous_on S (\x. norm (f x))" + using assms by (intro continuous_intros) + from continuous_attains_sup[OF `compact S` `S \ {}` this] + obtain M where M: "\x. x \ S \ norm (f x) \ M" + by auto + + show ?thesis + proof (rule integrable_bound) + show "integrable lborel (\x. indicator S x * M)" + using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left) + show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lborel" + using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact) + show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \ norm (indicator S x * M)" + by (auto split: split_indicator simp: abs_real_def dest!: M) + qed +qed simp + lemma borel_integrable_atLeastAtMost: fixes f :: "real \ real" assumes f: "\x. a \ x \ x \ b \ isCont f x" shows "integrable lborel (\x. f x * indicator {a .. b} x)" (is "integrable _ ?f") -proof cases - assume "a \ b" - - from isCont_Lb_Ub[OF `a \ b`, of f] f - obtain M L where - bounds: "\x. a \ x \ x \ b \ f x \ M" "\x. a \ x \ x \ b \ L \ f x" - by metis - - show ?thesis - proof (rule integrable_bound) - show "integrable lborel (\x. max \M\ \L\ * indicator {a..b} x)" - by (intro integrable_mult_right integrable_real_indicator) simp_all - show "AE x in lborel. norm (?f x) \ norm (max \M\ \L\ * indicator {a..b} x)" - proof (rule AE_I2) - fix x show "norm (?f x) \ norm (max \M\ \L\ * indicator {a..b} x)" - using bounds[of x] by (auto split: split_indicator) - qed - - let ?g = "\x. if x = a then f a else if x = b then f b else if x \ {a <..< b} then f x else 0" - from f have "continuous_on {a <..< b} f" - by (subst continuous_on_eq_continuous_at) auto - then have "?g \ borel_measurable borel" - using borel_measurable_continuous_on_open[of "{a <..< b }" f "\x. x" borel 0] - by (auto intro!: measurable_If[where P="\x. x = a"] measurable_If[where P="\x. x = b"]) - also have "?g = ?f" - using `a \ b` by (intro ext) (auto split: split_indicator) - finally show "?f \ borel_measurable lborel" - by simp - qed -qed simp +proof - + have "integrable lborel (\x. indicator {a .. b} x *\<^sub>R f x)" + proof (rule borel_integrable_compact) + from f show "continuous_on {a..b} f" + by (auto intro: continuous_at_imp_continuous_on) + qed simp + then show ?thesis + by (auto simp: mult_commute) +qed lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" diff -r f174712d0a84 -r 7b3146180291 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Fri May 30 15:56:30 2014 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Fri May 30 18:48:05 2014 +0200 @@ -2355,5 +2355,35 @@ f \ measurable M (restrict_space N \)" by (simp add: measurable_def space_restrict_space sets_restrict_space_iff Pi_Int[symmetric]) +lemma measurable_restrict_space_iff: + assumes \[simp, intro]: "\ \ space M \ sets M" "c \ space N" + shows "f \ measurable (restrict_space M \) N \ + (\x. if x \ \ then f x else c) \ measurable M N" (is "f \ measurable ?R N \ ?f \ measurable M N") + unfolding measurable_def +proof safe + fix x assume "f \ space ?R \ space N" "x \ space M" then show "?f x \ space N" + using `c\space N` by (auto simp: space_restrict_space) +next + fix x assume "?f \ space M \ space N" "x \ space ?R" then show "f x \ space N" + using `c\space N` by (auto simp: space_restrict_space Pi_iff) +next + fix X assume X: "X \ sets N" + assume *[THEN bspec]: "\y\sets N. f -` y \ space ?R \ sets ?R" + have "?f -` X \ space M = (f -` X \ (\ \ space M)) \ (if c \ X then (space M - (\ \ space M)) else {})" + by (auto split: split_if_asm) + also have "\ \ sets M" + using *[OF X] by (auto simp add: space_restrict_space sets_restrict_space_iff) + finally show "?f -` X \ space M \ sets M" . +next + assume *[THEN bspec]: "\y\sets N. ?f -` y \ space M \ sets M" + fix X :: "'b set" assume X: "X \ sets N" + have "f -` X \ (\ \ space M) = (?f -` X \ space M) \ (\ \ space M)" + by (auto simp: space_restrict_space) + also have "\ \ sets M" + using *[OF X] by auto + finally show "f -` X \ space ?R \ sets ?R" + by (auto simp add: sets_restrict_space_iff space_restrict_space) +qed + end