# HG changeset patch # User hoelzl # Date 1421330691 -3600 # Node ID fd5da2434be4c3fede604f2bfe5f8f64a7126718 # Parent b1e5d552e8cc6bb3b4186a658253a0672c2529e9 piecewise measurability using restrict_space; cleanup Borel_Space diff -r b1e5d552e8cc -r fd5da2434be4 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Wed Jan 14 17:04:19 2015 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Thu Jan 15 15:04:51 2015 +0100 @@ -202,21 +202,28 @@ qed auto qed (auto simp: eq intro: generate_topology.Basis) -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" +lemma borel_measurable_continuous_on_restrict: + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes f: "continuous_on A f" + shows "f \ borel_measurable (restrict_space borel A)" 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 + with f obtain T where "f -` S \ A = T \ A" "open T" + by (metis continuous_on_open_invariant) + then show "f -` S \ space (restrict_space borel A) \ sets (restrict_space borel A)" + by (force simp add: sets_restrict_space space_restrict_space) qed +lemma borel_measurable_continuous_on1: "continuous_on UNIV f \ f \ borel_measurable borel" + by (drule borel_measurable_continuous_on_restrict) simp + +lemma borel_measurable_continuous_on_if: + assumes [measurable]: "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" + by (rule measurable_piecewise_restrict[where + X="\b. if b then A else - A" and I=UNIV and f="\b x. if b then f x else g x"]) + (auto intro: f g borel_measurable_continuous_on_restrict) + lemma borel_measurable_continuous_countable_exceptions: fixes f :: "'a::t1_space \ 'b::topological_space" assumes X: "countable X" @@ -229,11 +236,6 @@ by (intro borel_measurable_continuous_on_if assms continuous_intros) qed auto -lemma borel_measurable_continuous_on1: - "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" @@ -636,7 +638,7 @@ fix x :: 'a assume "a < x \ i" with reals_Archimedean[of "x \ i - a"] obtain n where "a + 1 / real (Suc n) < x \ i" - by (auto simp: inverse_eq_divide field_simps) + by (auto simp: field_simps) then show "\n. a + 1 / real (Suc n) \ x \ i" by (blast intro: less_imp_le) next @@ -673,7 +675,7 @@ fix x::'a assume *: "x\i < a" with reals_Archimedean[of "a - x\i"] obtain n where "x \ i < a - 1 / (real (Suc n))" - by (auto simp: field_simps inverse_eq_divide) + by (auto simp: field_simps) then show "\n. x \ i \ a - 1 / (real (Suc n))" by (blast intro: less_imp_le) next @@ -683,7 +685,7 @@ finally show "x\i < a" . qed show "{x. x\i < a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) (auto intro: i) + by (intro sets.countable_UN) (auto intro: i) qed auto lemma borel_eq_halfspace_ge: @@ -693,7 +695,7 @@ fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" have *: "{x::'a. x\i < a} = space ?SIGMA - {x::'a. a \ x\i}" by auto show "{x. x\i < a} \ ?SIGMA" unfolding * - using i by (safe intro!: sets.compl_sets) auto + using i by (intro sets.compl_sets) auto qed auto lemma borel_eq_halfspace_greater: @@ -704,7 +706,7 @@ then have i: "i \ Basis" by auto have *: "{x::'a. x\i \ a} = space ?SIGMA - {x::'a. a < x\i}" by auto show "{x. x\i \ a} \ ?SIGMA" unfolding * - by (safe intro!: sets.compl_sets) (auto intro: i) + by (intro sets.compl_sets) (auto intro: i) qed auto lemma borel_eq_atMost: @@ -723,7 +725,7 @@ by (auto intro!: exI[of _ k]) qed show "{x. x\i \ a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) auto + by (intro sets.countable_UN) auto qed auto lemma borel_eq_greaterThan: @@ -748,7 +750,7 @@ qed finally show "{x. x\i \ a} \ ?SIGMA" apply (simp only:) - apply (safe intro!: sets.countable_UN sets.Diff) + apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top) done qed auto @@ -774,7 +776,7 @@ qed finally show "{x. a \ x\i} \ ?SIGMA" apply (simp only:) - apply (safe intro!: sets.countable_UN sets.Diff) + apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top ) done qed auto @@ -797,7 +799,7 @@ by (auto intro!: exI[of _ k]) qed show "{..a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) + by (intro sets.countable_UN) (auto intro!: sigma_sets_top) qed auto @@ -822,7 +824,7 @@ have "{..i::nat. {-real i ..< x})" by (auto simp: move_uminus real_arch_simple) then show "{y. y ?SIGMA" - by (auto intro: sigma_sets.intros simp: eucl_lessThan) + by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan) qed auto lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" @@ -831,15 +833,13 @@ fix x :: "'a set" assume "open x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect closed)" - by (rule sigma_sets.Compl) - (auto intro!: sigma_sets.Basic simp: `open x`) + by (force intro: sigma_sets.Compl simp: `open x`) finally show "x \ sigma_sets UNIV (Collect closed)" by simp next fix x :: "'a set" assume "closed x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect open)" - by (rule sigma_sets.Compl) - (auto intro!: sigma_sets.Basic simp: `closed x`) + by (force intro: sigma_sets.Compl simp: `closed x`) finally show "x \ sigma_sets UNIV (Collect open)" by simp qed simp_all @@ -1151,14 +1151,10 @@ 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 + apply (rule measurable_compose[OF f]) + apply (rule borel_measurable_continuous_countable_exceptions[of "{\, -\ }"]) + apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) + done lemma borel_measurable_ereal_cases: fixes f :: "'a \ ereal" @@ -1229,83 +1225,31 @@ finally show "f \ borel_measurable M" . qed simp_all -lemma borel_measurable_eq_atMost_ereal: - fixes f :: "'a \ ereal" - shows "f \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" -proof (intro iffI allI) - assume pos[rule_format]: "\a. f -` {..a} \ space M \ sets M" - show "f \ borel_measurable M" - unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le - proof (intro conjI allI) - fix a :: real - { fix x :: ereal assume *: "\i::nat. real i < x" - have "x = \" - proof (rule ereal_top) - fix B from reals_Archimedean2[of B] guess n .. - then have "ereal B < real n" by auto - with * show "B \ x" by (metis less_trans less_imp_le) - qed } - then have "f -` {\} \ space M = space M - (\i::nat. f -` {.. real i} \ space M)" - by (auto simp: not_le) - then show "f -` {\} \ space M \ sets M" using pos - by (auto simp del: UN_simps) - moreover - have "{-\::ereal} = {..-\}" by auto - then show "f -` {-\} \ space M \ sets M" using pos by auto - moreover have "{x\space M. f x \ ereal a} \ sets M" - using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute) - moreover have "{w \ space M. real (f w) \ a} = - (if a < 0 then {w \ space M. f w \ ereal a} - f -` {-\} \ space M - else {w \ space M. f w \ ereal a} \ (f -` {\} \ space M) \ (f -` {-\} \ space M))" (is "?l = ?r") - proof (intro set_eqI) fix x show "x \ ?l \ x \ ?r" by (cases "f x") auto qed - ultimately show "{w \ space M. real (f w) \ a} \ sets M" by auto - qed -qed (simp add: measurable_sets) +lemma borel_measurable_ereal_iff_Iio: + "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" + by (auto simp: borel_Iio measurable_iff_measure_of) + +lemma borel_measurable_ereal_iff_Ioi: + "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" + by (auto simp: borel_Ioi measurable_iff_measure_of) -lemma borel_measurable_eq_atLeast_ereal: - "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" -proof - assume pos: "\a. f -` {a..} \ space M \ sets M" - moreover have "\a. (\x. - f x) -` {..a} = f -` {-a ..}" - by (auto simp: ereal_uminus_le_reorder) - ultimately have "(\x. - f x) \ borel_measurable M" - unfolding borel_measurable_eq_atMost_ereal by auto - 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 +lemma vimage_sets_compl_iff: + "f -` A \ space M \ sets M \ f -` (- A) \ space M \ sets M" +proof - + { fix A assume "f -` A \ space M \ sets M" + moreover have "f -` (- A) \ space M = space M - f -` A \ space M" by auto + ultimately have "f -` (- A) \ space M \ sets M" by auto } + from this[of A] this[of "-A"] show ?thesis + by (metis double_complement) 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 borel_measurable_iff_Iic_ereal: + "(f::'a\ereal) \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" + unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp -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_iff_Ici_ereal: + "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" + unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp lemma borel_measurable_ereal2: fixes f g :: "'a \ ereal" @@ -1352,20 +1296,13 @@ fixes f :: "'c \ 'a \ ereal" 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 + using assms by (induction S rule: infinite_finite_induct) auto lemma borel_measurable_ereal_setprod[measurable (raw)]: fixes f :: "'c \ 'a \ ereal" 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 + using assms by (induction S rule: infinite_finite_induct) auto lemma [measurable (raw)]: fixes f :: "nat \ 'a \ ereal" diff -r b1e5d552e8cc -r fd5da2434be4 src/HOL/Probability/Measurable.thy --- a/src/HOL/Probability/Measurable.thy Wed Jan 14 17:04:19 2015 +0100 +++ b/src/HOL/Probability/Measurable.thy Thu Jan 15 15:04:51 2015 +0100 @@ -97,6 +97,7 @@ declare measurable_cong_sets[measurable_cong] declare sets_restrict_space_cong[measurable_cong] +declare sets_restrict_UNIV[measurable_cong] lemma predE[measurable (raw)]: "pred M P \ {x\space M. P x} \ sets M" diff -r b1e5d552e8cc -r fd5da2434be4 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Wed Jan 14 17:04:19 2015 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Thu Jan 15 15:04:51 2015 +0100 @@ -2440,6 +2440,9 @@ by simp qed +lemma sets_restrict_UNIV[simp]: "sets (restrict_space M UNIV) = sets M" + by (auto simp add: sets_restrict_space) + lemma sets_restrict_space_iff: "\ \ space M \ sets M \ A \ sets (restrict_space M \) \ (A \ \ \ A \ sets M)" proof (subst sets_restrict_space, safe) @@ -2525,5 +2528,40 @@ by (auto simp add: sets_restrict_space_iff space_restrict_space) qed +lemma measurableI: + "(\x. x \ space M \ f x \ space N) \ + (\A. A \ sets N \ f -` A \ space M \ sets M) \ + f \ measurable M N" + by (auto simp: measurable_def) + +lemma measurable_piecewise_restrict: + assumes I: "countable I" and X: "\i. i \ I \ X i \ sets M" "(\i\I. X i) = space M" + and meas: "\i. i \ I \ f i \ measurable (restrict_space M (X i)) N" + and eq: "\i x. i \ I \ x \ X i \ f i x = f' x" + shows "f' \ measurable M N" +proof (rule measurableI) + fix x assume "x \ space M" + then obtain i where "i \ I" "x \ X i" + using X by auto + then show "f' x \ space N" + using eq[of i x] meas[THEN measurable_space, of i x] `x \ space M` + by (auto simp: space_restrict_space) +next + fix A assume A: "A \ sets N" + have "f' -` A \ space M = {x \ space M. \i\I. x \ X i \ f i x \ A}" + by (auto simp: eq X(2)[symmetric]) + also have "\ \ sets M" + proof (intro sets.sets_Collect_countable_All'[OF _ I]) + fix i assume "i \ I" + then have "(f i -` A \ X i) \ (space M - X i) \ sets M" + using meas[THEN measurable_sets, OF `i \ I` A] X(1) + by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space) + also have "(f i -` A \ X i) \ (space M - X i) = {x \ space M. x \ X i \ f i x \ A}" + using `i \ I` by (auto simp: X(2)[symmetric]) + finally show "{x \ space M. x \ X i \ f i x \ A} \ sets M" . + qed + finally show "f' -` A \ space M \ sets M" . +qed + end