wenzelm@42150: (* Title: HOL/Probability/Borel_Space.thy hoelzl@42067: Author: Johannes Hölzl, TU München hoelzl@42067: Author: Armin Heller, TU München hoelzl@42067: *) hoelzl@38656: hoelzl@38656: header {*Borel spaces*} paulson@33533: hoelzl@40859: theory Borel_Space hoelzl@50387: imports hoelzl@50387: Measurable hoelzl@50387: "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis" paulson@33533: begin paulson@33533: hoelzl@56994: subsection {* Generic Borel spaces *} paulson@33533: hoelzl@47694: definition borel :: "'a::topological_space measure" where hoelzl@47694: "borel = sigma UNIV {S. open S}" paulson@33533: hoelzl@47694: abbreviation "borel_measurable M \ measurable M borel" paulson@33533: paulson@33533: lemma in_borel_measurable: paulson@33533: "f \ borel_measurable M \ hoelzl@47694: (\S \ sigma_sets UNIV {S. open S}. f -` S \ space M \ sets M)" hoelzl@40859: by (auto simp add: measurable_def borel_def) paulson@33533: hoelzl@40859: lemma in_borel_measurable_borel: hoelzl@38656: "f \ borel_measurable M \ hoelzl@40859: (\S \ sets borel. hoelzl@38656: f -` S \ space M \ sets M)" hoelzl@40859: by (auto simp add: measurable_def borel_def) paulson@33533: hoelzl@40859: lemma space_borel[simp]: "space borel = UNIV" hoelzl@40859: unfolding borel_def by auto hoelzl@38656: hoelzl@50002: lemma space_in_borel[measurable]: "UNIV \ sets borel" hoelzl@50002: unfolding borel_def by auto hoelzl@50002: hoelzl@57235: lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}" hoelzl@57235: unfolding borel_def by (rule sets_measure_of) simp hoelzl@57235: hoelzl@50387: lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \ {x. P x} \ sets borel" hoelzl@50002: unfolding borel_def pred_def by auto hoelzl@50002: hoelzl@50003: lemma borel_open[measurable (raw generic)]: hoelzl@40859: assumes "open A" shows "A \ sets borel" hoelzl@38656: proof - huffman@44537: have "A \ {S. open S}" unfolding mem_Collect_eq using assms . hoelzl@47694: thus ?thesis unfolding borel_def by auto paulson@33533: qed paulson@33533: hoelzl@50003: lemma borel_closed[measurable (raw generic)]: hoelzl@40859: assumes "closed A" shows "A \ sets borel" paulson@33533: proof - hoelzl@40859: have "space borel - (- A) \ sets borel" hoelzl@40859: using assms unfolding closed_def by (blast intro: borel_open) hoelzl@38656: thus ?thesis by simp paulson@33533: qed paulson@33533: hoelzl@50003: lemma borel_singleton[measurable]: hoelzl@50003: "A \ sets borel \ insert x A \ sets (borel :: 'a::t1_space measure)" immler@50244: unfolding insert_def by (rule sets.Un) auto hoelzl@50002: hoelzl@50003: lemma borel_comp[measurable]: "A \ sets borel \ - A \ sets borel" hoelzl@50002: unfolding Compl_eq_Diff_UNIV by simp hoelzl@41830: hoelzl@47694: lemma borel_measurable_vimage: hoelzl@38656: fixes f :: "'a \ 'x::t2_space" hoelzl@50002: assumes borel[measurable]: "f \ borel_measurable M" hoelzl@38656: shows "f -` {x} \ space M \ sets M" hoelzl@50002: by simp paulson@33533: hoelzl@47694: lemma borel_measurableI: hoelzl@38656: fixes f :: "'a \ 'x\topological_space" hoelzl@38656: assumes "\S. open S \ f -` S \ space M \ sets M" hoelzl@38656: shows "f \ borel_measurable M" hoelzl@40859: unfolding borel_def hoelzl@47694: proof (rule measurable_measure_of, simp_all) huffman@44537: fix S :: "'x set" assume "open S" thus "f -` S \ space M \ sets M" huffman@44537: using assms[of S] by simp hoelzl@40859: qed paulson@33533: hoelzl@50021: lemma borel_measurable_const: hoelzl@38656: "(\x. c) \ borel_measurable M" hoelzl@47694: by auto paulson@33533: hoelzl@50003: lemma borel_measurable_indicator: hoelzl@38656: assumes A: "A \ sets M" hoelzl@38656: shows "indicator A \ borel_measurable M" wenzelm@46905: unfolding indicator_def [abs_def] using A hoelzl@47694: by (auto intro!: measurable_If_set) paulson@33533: hoelzl@50096: lemma borel_measurable_count_space[measurable (raw)]: hoelzl@50096: "f \ borel_measurable (count_space S)" hoelzl@50096: unfolding measurable_def by auto hoelzl@50096: hoelzl@50096: lemma borel_measurable_indicator'[measurable (raw)]: hoelzl@50096: assumes [measurable]: "{x\space M. f x \ A x} \ sets M" hoelzl@50096: shows "(\x. indicator (A x) (f x)) \ borel_measurable M" hoelzl@50001: unfolding indicator_def[abs_def] hoelzl@50001: by (auto intro!: measurable_If) hoelzl@50001: hoelzl@47694: lemma borel_measurable_indicator_iff: hoelzl@40859: "(indicator A :: 'a \ 'x::{t1_space, zero_neq_one}) \ borel_measurable M \ A \ space M \ sets M" hoelzl@40859: (is "?I \ borel_measurable M \ _") hoelzl@40859: proof hoelzl@40859: assume "?I \ borel_measurable M" hoelzl@40859: then have "?I -` {1} \ space M \ sets M" hoelzl@40859: unfolding measurable_def by auto hoelzl@40859: also have "?I -` {1} \ space M = A \ space M" wenzelm@46905: unfolding indicator_def [abs_def] by auto hoelzl@40859: finally show "A \ space M \ sets M" . hoelzl@40859: next hoelzl@40859: assume "A \ space M \ sets M" hoelzl@40859: moreover have "?I \ borel_measurable M \ hoelzl@40859: (indicator (A \ space M) :: 'a \ 'x) \ borel_measurable M" hoelzl@40859: by (intro measurable_cong) (auto simp: indicator_def) hoelzl@40859: ultimately show "?I \ borel_measurable M" by auto hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma borel_measurable_subalgebra: hoelzl@41545: assumes "sets N \ sets M" "space N = space M" "f \ borel_measurable N" hoelzl@39092: shows "f \ borel_measurable M" hoelzl@39092: using assms unfolding measurable_def by auto hoelzl@39092: hoelzl@57137: lemma borel_measurable_restrict_space_iff_ereal: hoelzl@57137: fixes f :: "'a \ ereal" hoelzl@57137: assumes \[measurable, simp]: "\ \ space M \ sets M" hoelzl@57137: shows "f \ borel_measurable (restrict_space M \) \ hoelzl@57137: (\x. f x * indicator \ x) \ borel_measurable M" hoelzl@57138: by (subst measurable_restrict_space_iff) hoelzl@57138: (auto simp: indicator_def if_distrib[where f="\x. a * x" for a] cong del: if_cong) hoelzl@57137: hoelzl@57137: lemma borel_measurable_restrict_space_iff: hoelzl@57137: fixes f :: "'a \ 'b::real_normed_vector" hoelzl@57137: assumes \[measurable, simp]: "\ \ space M \ sets M" hoelzl@57137: shows "f \ borel_measurable (restrict_space M \) \ hoelzl@57137: (\x. indicator \ x *\<^sub>R f x) \ borel_measurable M" hoelzl@57138: by (subst measurable_restrict_space_iff) haftmann@57514: (auto simp: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps cong del: if_cong) hoelzl@57138: hoelzl@57138: lemma cbox_borel[measurable]: "cbox a b \ sets borel" hoelzl@57138: by (auto intro: borel_closed) hoelzl@57138: hoelzl@57447: lemma box_borel[measurable]: "box a b \ sets borel" hoelzl@57447: by (auto intro: borel_open) hoelzl@57447: hoelzl@57138: lemma borel_compact: "compact (A::'a::t2_space set) \ A \ sets borel" hoelzl@57138: by (auto intro: borel_closed dest!: compact_imp_closed) hoelzl@57137: hoelzl@57138: lemma borel_measurable_continuous_on_if: hoelzl@57138: assumes A: "A \ sets borel" and f: "continuous_on A f" and g: "continuous_on (- A) g" hoelzl@57138: shows "(\x. if x \ A then f x else g x) \ borel_measurable borel" hoelzl@57138: proof (rule borel_measurableI) hoelzl@57138: fix S :: "'b set" assume "open S" hoelzl@57138: have "(\x. if x \ A then f x else g x) -` S \ space borel = (f -` S \ A) \ (g -` S \ -A)" hoelzl@57138: by (auto split: split_if_asm) hoelzl@57138: moreover obtain A' where "f -` S \ A = A' \ A" "open A'" hoelzl@57138: using continuous_on_open_invariant[THEN iffD1, rule_format, OF f `open S`] by metis hoelzl@57138: moreover obtain B' where "g -` S \ -A = B' \ -A" "open B'" hoelzl@57138: using continuous_on_open_invariant[THEN iffD1, rule_format, OF g `open S`] by metis hoelzl@57138: ultimately show "(\x. if x \ A then f x else g x) -` S \ space borel \ sets borel" hoelzl@57138: using A by auto hoelzl@57137: qed hoelzl@57137: hoelzl@57275: lemma borel_measurable_continuous_countable_exceptions: hoelzl@57275: fixes f :: "'a::t1_space \ 'b::topological_space" hoelzl@57275: assumes X: "countable X" hoelzl@57275: assumes "continuous_on (- X) f" hoelzl@57275: shows "f \ borel_measurable borel" hoelzl@57275: proof (rule measurable_discrete_difference[OF _ X]) hoelzl@57275: have "X \ sets borel" hoelzl@57275: by (rule sets.countable[OF _ X]) auto hoelzl@57275: then show "(\x. if x \ X then undefined else f x) \ borel_measurable borel" hoelzl@57275: by (intro borel_measurable_continuous_on_if assms continuous_intros) hoelzl@57275: qed auto hoelzl@57275: hoelzl@50002: lemma borel_measurable_continuous_on1: hoelzl@57138: "continuous_on UNIV f \ f \ borel_measurable borel" hoelzl@57138: using borel_measurable_continuous_on_if[of UNIV f "\_. undefined"] hoelzl@57138: by (auto intro: continuous_on_const) hoelzl@57138: hoelzl@57138: lemma borel_measurable_continuous_on: hoelzl@57138: assumes f: "continuous_on UNIV f" and g: "g \ borel_measurable M" hoelzl@57138: shows "(\x. f (g x)) \ borel_measurable M" hoelzl@57138: using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def) hoelzl@57138: hoelzl@57138: lemma borel_measurable_continuous_on_open': hoelzl@57138: "continuous_on A f \ open A \ hoelzl@57138: (\x. if x \ A then f x else c) \ borel_measurable borel" hoelzl@57138: using borel_measurable_continuous_on_if[of A f "\_. c"] by (auto intro: continuous_on_const) hoelzl@57138: hoelzl@57138: lemma borel_measurable_continuous_on_open: hoelzl@57138: fixes f :: "'a::topological_space \ 'b::t1_space" hoelzl@57138: assumes cont: "continuous_on A f" "open A" hoelzl@57138: assumes g: "g \ borel_measurable M" hoelzl@57138: shows "(\x. if g x \ A then f (g x) else c) \ borel_measurable M" hoelzl@57138: using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c] hoelzl@57138: by (simp add: comp_def) hoelzl@57138: hoelzl@57138: lemma borel_measurable_continuous_on_indicator: hoelzl@57138: fixes f g :: "'a::topological_space \ 'b::real_normed_vector" hoelzl@57138: assumes A: "A \ sets borel" and f: "continuous_on A f" hoelzl@57138: shows "(\x. indicator A x *\<^sub>R f x) \ borel_measurable borel" hoelzl@57138: using borel_measurable_continuous_on_if[OF assms, of "\_. 0"] hoelzl@57138: by (simp add: indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] continuous_on_const hoelzl@57138: cong del: if_cong) hoelzl@50002: immler@50245: lemma borel_eq_countable_basis: immler@50245: fixes B::"'a::topological_space set set" immler@50245: assumes "countable B" immler@50245: assumes "topological_basis B" immler@50245: shows "borel = sigma UNIV B" immler@50087: unfolding borel_def immler@50087: proof (intro sigma_eqI sigma_sets_eqI, safe) immler@50245: interpret countable_basis using assms by unfold_locales immler@50245: fix X::"'a set" assume "open X" immler@50245: from open_countable_basisE[OF this] guess B' . note B' = this hoelzl@51683: then show "X \ sigma_sets UNIV B" hoelzl@51683: by (blast intro: sigma_sets_UNION `countable B` countable_subset) immler@50087: next immler@50245: fix b assume "b \ B" immler@50245: hence "open b" by (rule topological_basis_open[OF assms(2)]) immler@50245: thus "b \ sigma_sets UNIV (Collect open)" by auto immler@50087: qed simp_all immler@50087: hoelzl@50526: lemma borel_measurable_Pair[measurable (raw)]: hoelzl@50881: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" hoelzl@50526: assumes f[measurable]: "f \ borel_measurable M" hoelzl@50526: assumes g[measurable]: "g \ borel_measurable M" hoelzl@50526: shows "(\x. (f x, g x)) \ borel_measurable M" hoelzl@50526: proof (subst borel_eq_countable_basis) hoelzl@50526: let ?B = "SOME B::'b set set. countable B \ topological_basis B" hoelzl@50526: let ?C = "SOME B::'c set set. countable B \ topological_basis B" hoelzl@50526: let ?P = "(\(b, c). b \ c) ` (?B \ ?C)" hoelzl@50526: show "countable ?P" "topological_basis ?P" hoelzl@50526: by (auto intro!: countable_basis topological_basis_prod is_basis) hoelzl@38656: hoelzl@50526: show "(\x. (f x, g x)) \ measurable M (sigma UNIV ?P)" hoelzl@50526: proof (rule measurable_measure_of) hoelzl@50526: fix S assume "S \ ?P" hoelzl@50526: then obtain b c where "b \ ?B" "c \ ?C" and S: "S = b \ c" by auto hoelzl@50526: then have borel: "open b" "open c" hoelzl@50526: by (auto intro: is_basis topological_basis_open) hoelzl@50526: have "(\x. (f x, g x)) -` S \ space M = (f -` b \ space M) \ (g -` c \ space M)" hoelzl@50526: unfolding S by auto hoelzl@50526: also have "\ \ sets M" hoelzl@50526: using borel by simp hoelzl@50526: finally show "(\x. (f x, g x)) -` S \ space M \ sets M" . hoelzl@50526: qed auto hoelzl@39087: qed hoelzl@39087: hoelzl@49774: lemma borel_measurable_continuous_Pair: hoelzl@50881: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" hoelzl@50003: assumes [measurable]: "f \ borel_measurable M" hoelzl@50003: assumes [measurable]: "g \ borel_measurable M" hoelzl@49774: assumes H: "continuous_on UNIV (\x. H (fst x) (snd x))" hoelzl@49774: shows "(\x. H (f x) (g x)) \ borel_measurable M" hoelzl@49774: proof - hoelzl@49774: have eq: "(\x. H (f x) (g x)) = (\x. (\x. H (fst x) (snd x)) (f x, g x))" by auto hoelzl@49774: show ?thesis hoelzl@49774: unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto hoelzl@49774: qed hoelzl@49774: hoelzl@56994: subsection {* Borel spaces on euclidean spaces *} hoelzl@50526: hoelzl@50526: lemma borel_measurable_inner[measurable (raw)]: hoelzl@50881: fixes f g :: "'a \ 'b::{second_countable_topology, real_inner}" hoelzl@50526: assumes "f \ borel_measurable M" hoelzl@50526: assumes "g \ borel_measurable M" hoelzl@50526: shows "(\x. f x \ g x) \ borel_measurable M" hoelzl@50526: using assms hoelzl@56371: by (rule borel_measurable_continuous_Pair) (intro continuous_intros) hoelzl@50526: hoelzl@50526: lemma [measurable]: hoelzl@51683: fixes a b :: "'a\linorder_topology" hoelzl@50526: shows lessThan_borel: "{..< a} \ sets borel" hoelzl@50526: and greaterThan_borel: "{a <..} \ sets borel" hoelzl@50526: and greaterThanLessThan_borel: "{a<.. sets borel" hoelzl@50526: and atMost_borel: "{..a} \ sets borel" hoelzl@50526: and atLeast_borel: "{a..} \ sets borel" hoelzl@50526: and atLeastAtMost_borel: "{a..b} \ sets borel" hoelzl@50526: and greaterThanAtMost_borel: "{a<..b} \ sets borel" hoelzl@50526: and atLeastLessThan_borel: "{a.. sets borel" hoelzl@50526: unfolding greaterThanAtMost_def atLeastLessThan_def hoelzl@51683: by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan hoelzl@51683: closed_atMost closed_atLeast closed_atLeastAtMost)+ hoelzl@51683: immler@54775: notation immler@54775: eucl_less (infix " x \ b} = {x. a {..b}" immler@54775: and box_co: "{x. a \ x \ x {x. x ordered_euclidean_space" immler@54775: shows "{x. x sets borel" immler@54775: and "{x. a sets borel" hoelzl@51683: and "{..a} \ sets borel" hoelzl@51683: and "{a..} \ sets borel" hoelzl@51683: and "{a..b} \ sets borel" immler@54775: and "{x. a x \ b} \ sets borel" immler@54775: and "{x. a \ x \ x sets borel" immler@54775: unfolding box_oc box_co immler@54775: by (auto intro: borel_open borel_closed) hoelzl@50526: hoelzl@51683: lemma open_Collect_less: hoelzl@53216: fixes f g :: "'i::topological_space \ 'a :: {dense_linorder, linorder_topology}" hoelzl@51683: assumes "continuous_on UNIV f" hoelzl@51683: assumes "continuous_on UNIV g" hoelzl@51683: shows "open {x. f x < g x}" hoelzl@51683: proof - hoelzl@51683: have "open (\y. {x \ UNIV. f x \ {..< y}} \ {x \ UNIV. g x \ {y <..}})" (is "open ?X") hoelzl@51683: by (intro open_UN ballI open_Int continuous_open_preimage assms) auto hoelzl@51683: also have "?X = {x. f x < g x}" hoelzl@51683: by (auto intro: dense) hoelzl@51683: finally show ?thesis . hoelzl@51683: qed hoelzl@51683: hoelzl@51683: lemma closed_Collect_le: hoelzl@53216: fixes f g :: "'i::topological_space \ 'a :: {dense_linorder, linorder_topology}" hoelzl@51683: assumes f: "continuous_on UNIV f" hoelzl@51683: assumes g: "continuous_on UNIV g" hoelzl@51683: shows "closed {x. f x \ g x}" hoelzl@51683: using open_Collect_less[OF g f] unfolding not_less[symmetric] Collect_neg_eq open_closed . hoelzl@51683: hoelzl@50526: lemma borel_measurable_less[measurable]: hoelzl@53216: fixes f :: "'a \ 'b::{second_countable_topology, dense_linorder, linorder_topology}" hoelzl@51683: assumes "f \ borel_measurable M" hoelzl@51683: assumes "g \ borel_measurable M" hoelzl@50526: shows "{w \ space M. f w < g w} \ sets M" hoelzl@50526: proof - hoelzl@51683: have "{w \ space M. f w < g w} = (\x. (f x, g x)) -` {x. fst x < snd x} \ space M" hoelzl@51683: by auto hoelzl@51683: also have "\ \ sets M" hoelzl@51683: by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] hoelzl@56371: continuous_intros) hoelzl@51683: finally show ?thesis . hoelzl@50526: qed hoelzl@50526: hoelzl@50526: lemma hoelzl@53216: fixes f :: "'a \ 'b::{second_countable_topology, dense_linorder, linorder_topology}" hoelzl@50526: assumes f[measurable]: "f \ borel_measurable M" hoelzl@50526: assumes g[measurable]: "g \ borel_measurable M" hoelzl@50526: shows borel_measurable_le[measurable]: "{w \ space M. f w \ g w} \ sets M" hoelzl@50526: and borel_measurable_eq[measurable]: "{w \ space M. f w = g w} \ sets M" hoelzl@50526: and borel_measurable_neq: "{w \ space M. f w \ g w} \ sets M" hoelzl@50526: unfolding eq_iff not_less[symmetric] hoelzl@50526: by measurable hoelzl@50526: hoelzl@50526: lemma hoelzl@51683: fixes i :: "'a::{second_countable_topology, real_inner}" hoelzl@51683: shows hafspace_less_borel: "{x. a < x \ i} \ sets borel" hoelzl@51683: and hafspace_greater_borel: "{x. x \ i < a} \ sets borel" hoelzl@51683: and hafspace_less_eq_borel: "{x. a \ x \ i} \ sets borel" hoelzl@51683: and hafspace_greater_eq_borel: "{x. x \ i \ a} \ sets borel" hoelzl@50526: by simp_all hoelzl@50526: hoelzl@50526: subsection "Borel space equals sigma algebras over intervals" hoelzl@50526: hoelzl@50526: lemma borel_sigma_sets_subset: hoelzl@50526: "A \ sets borel \ sigma_sets UNIV A \ sets borel" hoelzl@50526: using sets.sigma_sets_subset[of A borel] by simp hoelzl@50526: hoelzl@50526: lemma borel_eq_sigmaI1: hoelzl@50526: fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" hoelzl@50526: assumes borel_eq: "borel = sigma UNIV X" hoelzl@50526: assumes X: "\x. x \ X \ x \ sets (sigma UNIV (F ` A))" hoelzl@50526: assumes F: "\i. i \ A \ F i \ sets borel" hoelzl@50526: shows "borel = sigma UNIV (F ` A)" hoelzl@50526: unfolding borel_def hoelzl@50526: proof (intro sigma_eqI antisym) hoelzl@50526: have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" hoelzl@50526: unfolding borel_def by simp hoelzl@50526: also have "\ = sigma_sets UNIV X" hoelzl@50526: unfolding borel_eq by simp hoelzl@50526: also have "\ \ sigma_sets UNIV (F`A)" hoelzl@50526: using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto hoelzl@50526: finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (F`A)" . hoelzl@50526: show "sigma_sets UNIV (F`A) \ sigma_sets UNIV {S. open S}" hoelzl@50526: unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_sigmaI2: hoelzl@50526: fixes F :: "'i \ 'j \ 'a::topological_space set" hoelzl@50526: and G :: "'l \ 'k \ 'a::topological_space set" hoelzl@50526: assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`B)" hoelzl@50526: assumes X: "\i j. (i, j) \ B \ G i j \ sets (sigma UNIV ((\(i, j). F i j) ` A))" hoelzl@50526: assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" hoelzl@50526: shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" hoelzl@50526: using assms hoelzl@50526: by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` B" and F="(\(i, j). F i j)"]) auto hoelzl@50526: hoelzl@50526: lemma borel_eq_sigmaI3: hoelzl@50526: fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" hoelzl@50526: assumes borel_eq: "borel = sigma UNIV X" hoelzl@50526: assumes X: "\x. x \ X \ x \ sets (sigma UNIV ((\(i, j). F i j) ` A))" hoelzl@50526: assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" hoelzl@50526: shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" hoelzl@50526: using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto hoelzl@50526: hoelzl@50526: lemma borel_eq_sigmaI4: hoelzl@50526: fixes F :: "'i \ 'a::topological_space set" hoelzl@50526: and G :: "'l \ 'k \ 'a::topological_space set" hoelzl@50526: assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`A)" hoelzl@50526: assumes X: "\i j. (i, j) \ A \ G i j \ sets (sigma UNIV (range F))" hoelzl@50526: assumes F: "\i. F i \ sets borel" hoelzl@50526: shows "borel = sigma UNIV (range F)" hoelzl@50526: using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` A" and F=F]) auto hoelzl@50526: hoelzl@50526: lemma borel_eq_sigmaI5: hoelzl@50526: fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" hoelzl@50526: assumes borel_eq: "borel = sigma UNIV (range G)" hoelzl@50526: assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" hoelzl@50526: assumes F: "\i j. F i j \ sets borel" hoelzl@50526: shows "borel = sigma UNIV (range (\(i, j). F i j))" hoelzl@50526: using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto hoelzl@50526: hoelzl@50526: lemma borel_eq_box: hoelzl@50526: "borel = sigma UNIV (range (\ (a, b). box a b :: 'a \ euclidean_space set))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI1[OF borel_def]) hoelzl@50526: fix M :: "'a set" assume "M \ {S. open S}" hoelzl@50526: then have "open M" by simp hoelzl@50526: show "M \ ?SIGMA" hoelzl@50526: apply (subst open_UNION_box[OF `open M`]) hoelzl@50526: apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) hoelzl@50526: apply (auto intro: countable_rat) hoelzl@50526: done hoelzl@50526: qed (auto simp: box_def) hoelzl@50526: hoelzl@50526: lemma halfspace_gt_in_halfspace: hoelzl@50526: assumes i: "i \ A" hoelzl@50526: shows "{x\'a. a < x \ i} \ hoelzl@50526: sigma_sets UNIV ((\ (a, i). {x\'a\euclidean_space. x \ i < a}) ` (UNIV \ A))" hoelzl@50526: (is "?set \ ?SIGMA") hoelzl@50526: proof - hoelzl@50526: interpret sigma_algebra UNIV ?SIGMA hoelzl@50526: by (intro sigma_algebra_sigma_sets) simp_all hoelzl@50526: have *: "?set = (\n. UNIV - {x\'a. x \ i < a + 1 / real (Suc n)})" hoelzl@50526: proof (safe, simp_all add: not_less) hoelzl@50526: fix x :: 'a assume "a < x \ i" hoelzl@50526: with reals_Archimedean[of "x \ i - a"] hoelzl@50526: obtain n where "a + 1 / real (Suc n) < x \ i" hoelzl@50526: by (auto simp: inverse_eq_divide field_simps) hoelzl@50526: then show "\n. a + 1 / real (Suc n) \ x \ i" hoelzl@50526: by (blast intro: less_imp_le) hoelzl@50526: next hoelzl@50526: fix x n hoelzl@50526: have "a < a + 1 / real (Suc n)" by auto hoelzl@50526: also assume "\ \ x" hoelzl@50526: finally show "a < x" . hoelzl@50526: qed hoelzl@50526: show "?set \ ?SIGMA" unfolding * hoelzl@50526: by (auto del: Diff intro!: Diff i) hoelzl@50526: qed hoelzl@50526: hoelzl@50526: lemma borel_eq_halfspace_less: hoelzl@50526: "borel = sigma UNIV ((\(a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ Basis))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI2[OF borel_eq_box]) hoelzl@50526: fix a b :: 'a hoelzl@50526: have "box a b = {x\space ?SIGMA. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" hoelzl@50526: by (auto simp: box_def) hoelzl@50526: also have "\ \ sets ?SIGMA" hoelzl@50526: by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const) hoelzl@50526: (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat) hoelzl@50526: finally show "box a b \ sets ?SIGMA" . hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_halfspace_le: hoelzl@50526: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i \ a}) ` (UNIV \ Basis))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) hoelzl@50526: fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" hoelzl@50526: then have i: "i \ Basis" by auto hoelzl@50526: have *: "{x::'a. x\i < a} = (\n. {x. x\i \ a - 1/real (Suc n)})" hoelzl@50526: proof (safe, simp_all) hoelzl@50526: fix x::'a assume *: "x\i < a" hoelzl@50526: with reals_Archimedean[of "a - x\i"] hoelzl@50526: obtain n where "x \ i < a - 1 / (real (Suc n))" hoelzl@50526: by (auto simp: field_simps inverse_eq_divide) hoelzl@50526: then show "\n. x \ i \ a - 1 / (real (Suc n))" hoelzl@50526: by (blast intro: less_imp_le) hoelzl@50526: next hoelzl@50526: fix x::'a and n hoelzl@50526: assume "x\i \ a - 1 / real (Suc n)" hoelzl@50526: also have "\ < a" by auto hoelzl@50526: finally show "x\i < a" . hoelzl@50526: qed hoelzl@50526: show "{x. x\i < a} \ ?SIGMA" unfolding * hoelzl@50526: by (safe intro!: sets.countable_UN) (auto intro: i) hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_halfspace_ge: hoelzl@50526: "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) hoelzl@50526: fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" hoelzl@50526: have *: "{x::'a. x\i < a} = space ?SIGMA - {x::'a. a \ x\i}" by auto hoelzl@50526: show "{x. x\i < a} \ ?SIGMA" unfolding * hoelzl@50526: using i by (safe intro!: sets.compl_sets) auto hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_halfspace_greater: hoelzl@50526: "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a < x \ i}) ` (UNIV \ Basis))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) hoelzl@50526: fix a :: real and i :: 'a assume "(a, i) \ (UNIV \ Basis)" hoelzl@50526: then have i: "i \ Basis" by auto hoelzl@50526: have *: "{x::'a. x\i \ a} = space ?SIGMA - {x::'a. a < x\i}" by auto hoelzl@50526: show "{x. x\i \ a} \ ?SIGMA" unfolding * hoelzl@50526: by (safe intro!: sets.compl_sets) (auto intro: i) hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_atMost: hoelzl@50526: "borel = sigma UNIV (range (\a. {..a\'a\ordered_euclidean_space}))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) hoelzl@50526: fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" hoelzl@50526: then have "i \ Basis" by auto hoelzl@50526: then have *: "{x::'a. x\i \ a} = (\k::nat. {.. (\n\Basis. (if n = i then a else real k)*\<^sub>R n)})" hoelzl@50526: proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) hoelzl@50526: fix x :: 'a hoelzl@50526: from real_arch_simple[of "Max ((\i. x\i)`Basis)"] guess k::nat .. hoelzl@50526: then have "\i. i \ Basis \ x\i \ real k" hoelzl@50526: by (subst (asm) Max_le_iff) auto hoelzl@50526: then show "\k::nat. \ia\Basis. ia \ i \ x \ ia \ real k" hoelzl@50526: by (auto intro!: exI[of _ k]) hoelzl@50526: qed hoelzl@50526: show "{x. x\i \ a} \ ?SIGMA" unfolding * hoelzl@50526: by (safe intro!: sets.countable_UN) auto hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_greaterThan: immler@54775: "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {x. a UNIV \ Basis" hoelzl@50526: then have i: "i \ Basis" by auto hoelzl@50526: have "{x::'a. x\i \ a} = UNIV - {x::'a. a < x\i}" by auto hoelzl@50526: also have *: "{x::'a. a < x\i} = immler@54775: (\k::nat. {x. (\n\Basis. (if n = i then a else -real k) *\<^sub>R n) i. -x\i)`Basis)"] hoelzl@50526: guess k::nat .. note k = this hoelzl@50526: { fix i :: 'a assume "i \ Basis" hoelzl@50526: then have "-x\i < real k" hoelzl@50526: using k by (subst (asm) Max_less_iff) auto hoelzl@50526: then have "- real k < x\i" by simp } hoelzl@50526: then show "\k::nat. \ia\Basis. ia \ i \ -real k < x \ ia" hoelzl@50526: by (auto intro!: exI[of _ k]) hoelzl@50526: qed hoelzl@50526: finally show "{x. x\i \ a} \ ?SIGMA" hoelzl@50526: apply (simp only:) hoelzl@50526: apply (safe intro!: sets.countable_UN sets.Diff) hoelzl@50526: apply (auto intro: sigma_sets_top) hoelzl@50526: done hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_lessThan: immler@54775: "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {x. x UNIV \ Basis" hoelzl@50526: then have i: "i \ Basis" by auto hoelzl@50526: have "{x::'a. a \ x\i} = UNIV - {x::'a. x\i < a}" by auto immler@54775: also have *: "{x::'a. x\i < a} = (\k::nat. {x. x n\Basis. (if n = i then a else real k) *\<^sub>R n)})" using `i\ Basis` immler@54775: proof (safe, simp_all add: eucl_less_def split: split_if_asm) hoelzl@50526: fix x :: 'a hoelzl@50526: from reals_Archimedean2[of "Max ((\i. x\i)`Basis)"] hoelzl@50526: guess k::nat .. note k = this hoelzl@50526: { fix i :: 'a assume "i \ Basis" hoelzl@50526: then have "x\i < real k" hoelzl@50526: using k by (subst (asm) Max_less_iff) auto hoelzl@50526: then have "x\i < real k" by simp } hoelzl@50526: then show "\k::nat. \ia\Basis. ia \ i \ x \ ia < real k" hoelzl@50526: by (auto intro!: exI[of _ k]) hoelzl@50526: qed hoelzl@50526: finally show "{x. a \ x\i} \ ?SIGMA" hoelzl@50526: apply (simp only:) hoelzl@50526: apply (safe intro!: sets.countable_UN sets.Diff) immler@54775: apply (auto intro: sigma_sets_top ) hoelzl@50526: done hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_atLeastAtMost: hoelzl@50526: "borel = sigma UNIV (range (\(a,b). {a..b} \'a\ordered_euclidean_space set))" hoelzl@50526: (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) hoelzl@50526: fix a::'a hoelzl@50526: have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" hoelzl@50526: proof (safe, simp_all add: eucl_le[where 'a='a]) hoelzl@50526: fix x :: 'a hoelzl@50526: from real_arch_simple[of "Max ((\i. - x\i)`Basis)"] hoelzl@50526: guess k::nat .. note k = this hoelzl@50526: { fix i :: 'a assume "i \ Basis" hoelzl@50526: with k have "- x\i \ real k" hoelzl@50526: by (subst (asm) Max_le_iff) (auto simp: field_simps) hoelzl@50526: then have "- real k \ x\i" by simp } hoelzl@50526: then show "\n::nat. \i\Basis. - real n \ x \ i" hoelzl@50526: by (auto intro!: exI[of _ k]) hoelzl@50526: qed hoelzl@50526: show "{..a} \ ?SIGMA" unfolding * hoelzl@50526: by (safe intro!: sets.countable_UN) hoelzl@50526: (auto intro!: sigma_sets_top) hoelzl@50526: qed auto hoelzl@50526: hoelzl@57447: lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\(a, b). {a <.. b::real}))" hoelzl@57447: proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) hoelzl@57447: fix i :: real hoelzl@57447: have "{..i} = (\j::nat. {-j <.. i})" hoelzl@57447: by (auto simp: minus_less_iff reals_Archimedean2) hoelzl@57447: also have "\ \ sets (sigma UNIV (range (\(i, j). {i<..j})))" hoelzl@57447: by (intro sets.countable_nat_UN) auto hoelzl@57447: finally show "{..i} \ sets (sigma UNIV (range (\(i, j). {i<..j})))" . hoelzl@57447: qed simp hoelzl@57447: immler@54775: lemma eucl_lessThan: "{x::real. x (a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") hoelzl@50526: proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) hoelzl@50526: have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto hoelzl@50526: fix x :: real hoelzl@50526: have "{..i::nat. {-real i ..< x})" hoelzl@50526: by (auto simp: move_uminus real_arch_simple) immler@54775: then show "{y. y ?SIGMA" immler@54775: by (auto intro: sigma_sets.intros simp: eucl_lessThan) hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" hoelzl@50526: unfolding borel_def hoelzl@50526: proof (intro sigma_eqI sigma_sets_eqI, safe) hoelzl@50526: fix x :: "'a set" assume "open x" hoelzl@50526: hence "x = UNIV - (UNIV - x)" by auto hoelzl@50526: also have "\ \ sigma_sets UNIV (Collect closed)" hoelzl@50526: by (rule sigma_sets.Compl) hoelzl@50526: (auto intro!: sigma_sets.Basic simp: `open x`) hoelzl@50526: finally show "x \ sigma_sets UNIV (Collect closed)" by simp hoelzl@50526: next hoelzl@50526: fix x :: "'a set" assume "closed x" hoelzl@50526: hence "x = UNIV - (UNIV - x)" by auto hoelzl@50526: also have "\ \ sigma_sets UNIV (Collect open)" hoelzl@50526: by (rule sigma_sets.Compl) hoelzl@50526: (auto intro!: sigma_sets.Basic simp: `closed x`) hoelzl@50526: finally show "x \ sigma_sets UNIV (Collect open)" by simp hoelzl@50526: qed simp_all hoelzl@50526: hoelzl@50526: lemma borel_measurable_halfspacesI: hoelzl@50526: fixes f :: "'a \ 'c\euclidean_space" hoelzl@50526: assumes F: "borel = sigma UNIV (F ` (UNIV \ Basis))" hoelzl@50526: and S_eq: "\a i. S a i = f -` F (a,i) \ space M" hoelzl@50526: shows "f \ borel_measurable M = (\i\Basis. \a::real. S a i \ sets M)" hoelzl@50526: proof safe hoelzl@50526: fix a :: real and i :: 'b assume i: "i \ Basis" and f: "f \ borel_measurable M" hoelzl@50526: then show "S a i \ sets M" unfolding assms hoelzl@50526: by (auto intro!: measurable_sets simp: assms(1)) hoelzl@50526: next hoelzl@50526: assume a: "\i\Basis. \a. S a i \ sets M" hoelzl@50526: then show "f \ borel_measurable M" hoelzl@50526: by (auto intro!: measurable_measure_of simp: S_eq F) hoelzl@50526: qed hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_halfspace_le: hoelzl@50526: fixes f :: "'a \ 'c\euclidean_space" hoelzl@50526: shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. f w \ i \ a} \ sets M)" hoelzl@50526: by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_halfspace_less: hoelzl@50526: fixes f :: "'a \ 'c\euclidean_space" hoelzl@50526: shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. f w \ i < a} \ sets M)" hoelzl@50526: by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_halfspace_ge: hoelzl@50526: fixes f :: "'a \ 'c\euclidean_space" hoelzl@50526: shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. a \ f w \ i} \ sets M)" hoelzl@50526: by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_halfspace_greater: hoelzl@50526: fixes f :: "'a \ 'c\euclidean_space" hoelzl@50526: shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. a < f w \ i} \ sets M)" hoelzl@50526: by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_le: hoelzl@50526: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" hoelzl@50526: using borel_measurable_iff_halfspace_le[where 'c=real] by simp hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_less: hoelzl@50526: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" hoelzl@50526: using borel_measurable_iff_halfspace_less[where 'c=real] by simp hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_ge: hoelzl@50526: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" hoelzl@50526: using borel_measurable_iff_halfspace_ge[where 'c=real] hoelzl@50526: by simp hoelzl@50526: hoelzl@50526: lemma borel_measurable_iff_greater: hoelzl@50526: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" hoelzl@50526: using borel_measurable_iff_halfspace_greater[where 'c=real] by simp hoelzl@50526: hoelzl@50526: lemma borel_measurable_euclidean_space: hoelzl@50526: fixes f :: "'a \ 'c::euclidean_space" hoelzl@50526: shows "f \ borel_measurable M \ (\i\Basis. (\x. f x \ i) \ borel_measurable M)" hoelzl@50526: proof safe hoelzl@50526: assume f: "\i\Basis. (\x. f x \ i) \ borel_measurable M" hoelzl@50526: then show "f \ borel_measurable M" hoelzl@50526: by (subst borel_measurable_iff_halfspace_le) auto hoelzl@50526: qed auto hoelzl@50526: hoelzl@50526: subsection "Borel measurable operators" hoelzl@50526: hoelzl@56993: lemma borel_measurable_norm[measurable]: "norm \ borel_measurable borel" hoelzl@56993: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@56993: hoelzl@57275: lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \ 'a) \ borel_measurable borel" hoelzl@57275: by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"]) hoelzl@57275: (auto intro!: continuous_on_sgn continuous_on_id) hoelzl@57275: hoelzl@50526: lemma borel_measurable_uminus[measurable (raw)]: hoelzl@51683: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" hoelzl@50526: assumes g: "g \ borel_measurable M" hoelzl@50526: shows "(\x. - g x) \ borel_measurable M" hoelzl@56371: by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros) hoelzl@50526: hoelzl@50003: lemma borel_measurable_add[measurable (raw)]: hoelzl@51683: fixes f g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes g: "g \ borel_measurable M" hoelzl@49774: shows "(\x. f x + g x) \ borel_measurable M" hoelzl@56371: using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) hoelzl@49774: hoelzl@50003: lemma borel_measurable_setsum[measurable (raw)]: hoelzl@51683: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, real_normed_vector}" hoelzl@49774: assumes "\i. i \ S \ f i \ borel_measurable M" hoelzl@49774: shows "(\x. \i\S. f i x) \ borel_measurable M" hoelzl@49774: proof cases hoelzl@49774: assume "finite S" hoelzl@49774: thus ?thesis using assms by induct auto hoelzl@49774: qed simp hoelzl@49774: hoelzl@50003: lemma borel_measurable_diff[measurable (raw)]: hoelzl@51683: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_vector}" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes g: "g \ borel_measurable M" hoelzl@49774: shows "(\x. f x - g x) \ borel_measurable M" haftmann@54230: using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) hoelzl@49774: hoelzl@50003: lemma borel_measurable_times[measurable (raw)]: hoelzl@51683: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_algebra}" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes g: "g \ borel_measurable M" hoelzl@49774: shows "(\x. f x * g x) \ borel_measurable M" hoelzl@56371: using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) hoelzl@51683: hoelzl@51683: lemma borel_measurable_setprod[measurable (raw)]: hoelzl@51683: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, real_normed_field}" hoelzl@51683: assumes "\i. i \ S \ f i \ borel_measurable M" hoelzl@51683: shows "(\x. \i\S. f i x) \ borel_measurable M" hoelzl@51683: proof cases hoelzl@51683: assume "finite S" hoelzl@51683: thus ?thesis using assms by induct auto hoelzl@51683: qed simp hoelzl@49774: hoelzl@50003: lemma borel_measurable_dist[measurable (raw)]: hoelzl@51683: fixes g f :: "'a \ 'b::{second_countable_topology, metric_space}" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes g: "g \ borel_measurable M" hoelzl@49774: shows "(\x. dist (f x) (g x)) \ borel_measurable M" hoelzl@56371: using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) hoelzl@49774: hoelzl@50002: lemma borel_measurable_scaleR[measurable (raw)]: hoelzl@51683: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" hoelzl@50002: assumes f: "f \ borel_measurable M" hoelzl@50002: assumes g: "g \ borel_measurable M" hoelzl@50002: shows "(\x. f x *\<^sub>R g x) \ borel_measurable M" hoelzl@56371: using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) hoelzl@50002: hoelzl@47694: lemma affine_borel_measurable_vector: hoelzl@38656: fixes f :: "'a \ 'x::real_normed_vector" hoelzl@38656: assumes "f \ borel_measurable M" hoelzl@38656: shows "(\x. a + b *\<^sub>R f x) \ borel_measurable M" hoelzl@38656: proof (rule borel_measurableI) hoelzl@38656: fix S :: "'x set" assume "open S" hoelzl@38656: show "(\x. a + b *\<^sub>R f x) -` S \ space M \ sets M" hoelzl@38656: proof cases hoelzl@38656: assume "b \ 0" huffman@44537: with `open S` have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") haftmann@54230: using open_affinity [of S "inverse b" "- a /\<^sub>R b"] haftmann@54230: by (auto simp: algebra_simps) hoelzl@47694: hence "?S \ sets borel" by auto hoelzl@38656: moreover hoelzl@38656: from `b \ 0` have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" hoelzl@38656: apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) hoelzl@40859: ultimately show ?thesis using assms unfolding in_borel_measurable_borel hoelzl@38656: by auto hoelzl@38656: qed simp hoelzl@38656: qed hoelzl@38656: hoelzl@50002: lemma borel_measurable_const_scaleR[measurable (raw)]: hoelzl@50002: "f \ borel_measurable M \ (\x. b *\<^sub>R f x ::'a::real_normed_vector) \ borel_measurable M" hoelzl@50002: using affine_borel_measurable_vector[of f M 0 b] by simp hoelzl@38656: hoelzl@50002: lemma borel_measurable_const_add[measurable (raw)]: hoelzl@50002: "f \ borel_measurable M \ (\x. a + f x ::'a::real_normed_vector) \ borel_measurable M" hoelzl@50002: using affine_borel_measurable_vector[of f M a 1] by simp hoelzl@50002: hoelzl@50003: lemma borel_measurable_inverse[measurable (raw)]: hoelzl@57275: fixes f :: "'a \ 'b::real_normed_div_algebra" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@35692: shows "(\x. inverse (f x)) \ borel_measurable M" hoelzl@57275: apply (rule measurable_compose[OF f]) hoelzl@57275: apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) hoelzl@57275: apply (auto intro!: continuous_on_inverse continuous_on_id) hoelzl@57275: done hoelzl@35692: hoelzl@50003: lemma borel_measurable_divide[measurable (raw)]: hoelzl@51683: "f \ borel_measurable M \ g \ borel_measurable M \ hoelzl@57275: (\x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \ borel_measurable M" hoelzl@57275: by (simp add: divide_inverse) hoelzl@38656: hoelzl@50003: lemma borel_measurable_max[measurable (raw)]: hoelzl@53216: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. max (g x) (f x) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \ borel_measurable M" hoelzl@50003: by (simp add: max_def) hoelzl@38656: hoelzl@50003: lemma borel_measurable_min[measurable (raw)]: hoelzl@53216: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. min (g x) (f x) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \ borel_measurable M" hoelzl@50003: by (simp add: min_def) hoelzl@38656: hoelzl@57235: lemma borel_measurable_Min[measurable (raw)]: hoelzl@57235: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Min ((\i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \ borel_measurable M" hoelzl@57235: proof (induct I rule: finite_induct) hoelzl@57235: case (insert i I) then show ?case hoelzl@57235: by (cases "I = {}") auto hoelzl@57235: qed auto hoelzl@57235: hoelzl@57235: lemma borel_measurable_Max[measurable (raw)]: hoelzl@57235: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Max ((\i. f i x)`I) :: 'b::{second_countable_topology, dense_linorder, linorder_topology}) \ borel_measurable M" hoelzl@57235: proof (induct I rule: finite_induct) hoelzl@57235: case (insert i I) then show ?case hoelzl@57235: by (cases "I = {}") auto hoelzl@57235: qed auto hoelzl@57235: hoelzl@50003: lemma borel_measurable_abs[measurable (raw)]: hoelzl@50003: "f \ borel_measurable M \ (\x. \f x :: real\) \ borel_measurable M" hoelzl@50003: unfolding abs_real_def by simp hoelzl@38656: hoelzl@50003: lemma borel_measurable_nth[measurable (raw)]: hoelzl@41026: "(\x::real^'n. x $ i) \ borel_measurable borel" hoelzl@50526: by (simp add: cart_eq_inner_axis) hoelzl@41026: hoelzl@47694: lemma convex_measurable: hoelzl@51683: fixes A :: "'a :: ordered_euclidean_space set" hoelzl@51683: assumes X: "X \ borel_measurable M" "X ` space M \ A" "open A" hoelzl@51683: assumes q: "convex_on A q" hoelzl@49774: shows "(\x. q (X x)) \ borel_measurable M" hoelzl@42990: proof - hoelzl@51683: have "(\x. if X x \ A then q (X x) else 0) \ borel_measurable M" (is "?qX") hoelzl@49774: proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)]) hoelzl@51683: show "open A" by fact hoelzl@51683: from this q show "continuous_on A q" hoelzl@42990: by (rule convex_on_continuous) hoelzl@41830: qed hoelzl@50002: also have "?qX \ (\x. q (X x)) \ borel_measurable M" hoelzl@42990: using X by (intro measurable_cong) auto hoelzl@50002: finally show ?thesis . hoelzl@41830: qed hoelzl@41830: hoelzl@50003: lemma borel_measurable_ln[measurable (raw)]: hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: shows "(\x. ln (f x)) \ borel_measurable M" hoelzl@57275: apply (rule measurable_compose[OF f]) hoelzl@57275: apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) hoelzl@57275: apply (auto intro!: continuous_on_ln continuous_on_id) hoelzl@57275: done hoelzl@41830: hoelzl@50003: lemma borel_measurable_log[measurable (raw)]: hoelzl@50002: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. log (g x) (f x)) \ borel_measurable M" hoelzl@49774: unfolding log_def by auto hoelzl@41830: hoelzl@50419: lemma borel_measurable_exp[measurable]: "exp \ borel_measurable borel" hoelzl@51478: by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) hoelzl@50419: hoelzl@50002: lemma measurable_real_floor[measurable]: hoelzl@50002: "(floor :: real \ int) \ measurable borel (count_space UNIV)" hoelzl@47761: proof - hoelzl@50002: have "\a x. \x\ = a \ (real a \ x \ x < real (a + 1))" hoelzl@50002: by (auto intro: floor_eq2) hoelzl@50002: then show ?thesis hoelzl@50002: by (auto simp: vimage_def measurable_count_space_eq2_countable) hoelzl@47761: qed hoelzl@47761: hoelzl@50002: lemma measurable_real_natfloor[measurable]: hoelzl@50002: "(natfloor :: real \ nat) \ measurable borel (count_space UNIV)" hoelzl@50002: by (simp add: natfloor_def[abs_def]) hoelzl@50002: hoelzl@50002: lemma measurable_real_ceiling[measurable]: hoelzl@50002: "(ceiling :: real \ int) \ measurable borel (count_space UNIV)" hoelzl@50002: unfolding ceiling_def[abs_def] by simp hoelzl@50002: hoelzl@50002: lemma borel_measurable_real_floor: "(\x::real. real \x\) \ borel_measurable borel" hoelzl@50002: by simp hoelzl@50002: hoelzl@50003: lemma borel_measurable_real_natfloor: hoelzl@50002: "f \ borel_measurable M \ (\x. real (natfloor (f x))) \ borel_measurable M" hoelzl@50002: by simp hoelzl@50002: hoelzl@57235: lemma borel_measurable_root [measurable]: "(root n) \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_sqrt [measurable]: "sqrt \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_power [measurable (raw)]: hoelzl@57235: fixes f :: "_ \ 'b::{power,real_normed_algebra}" hoelzl@57235: assumes f: "f \ borel_measurable M" hoelzl@57235: shows "(\x. (f x) ^ n) \ borel_measurable M" hoelzl@57235: by (intro borel_measurable_continuous_on [OF _ f] continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_Re [measurable]: "Re \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_Im [measurable]: "Im \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_of_real [measurable]: "(of_real :: _ \ (_::real_normed_algebra)) \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_sin [measurable]: "sin \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_cos [measurable]: "cos \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57235: lemma borel_measurable_arctan [measurable]: "arctan \ borel_measurable borel" hoelzl@57235: by (intro borel_measurable_continuous_on1 continuous_intros) hoelzl@57235: hoelzl@57259: lemma borel_measurable_complex_iff: hoelzl@57259: "f \ borel_measurable M \ hoelzl@57259: (\x. Re (f x)) \ borel_measurable M \ (\x. Im (f x)) \ borel_measurable M" hoelzl@57259: apply auto hoelzl@57259: apply (subst fun_complex_eq) hoelzl@57259: apply (intro borel_measurable_add) hoelzl@57259: apply auto hoelzl@57259: done hoelzl@57259: hoelzl@41981: subsection "Borel space on the extended reals" hoelzl@41981: hoelzl@50003: lemma borel_measurable_ereal[measurable (raw)]: hoelzl@43920: assumes f: "f \ borel_measurable M" shows "(\x. ereal (f x)) \ borel_measurable M" hoelzl@49774: using continuous_on_ereal f by (rule borel_measurable_continuous_on) hoelzl@41981: hoelzl@50003: lemma borel_measurable_real_of_ereal[measurable (raw)]: hoelzl@49774: fixes f :: "'a \ ereal" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: shows "(\x. real (f x)) \ borel_measurable M" hoelzl@49774: proof - hoelzl@49774: have "(\x. if f x \ UNIV - { \, - \ } then real (f x) else 0) \ borel_measurable M" hoelzl@49774: using continuous_on_real hoelzl@49774: by (rule borel_measurable_continuous_on_open[OF _ _ f]) auto hoelzl@49774: also have "(\x. if f x \ UNIV - { \, - \ } then real (f x) else 0) = (\x. real (f x))" hoelzl@49774: by auto hoelzl@49774: finally show ?thesis . hoelzl@49774: qed hoelzl@49774: hoelzl@49774: lemma borel_measurable_ereal_cases: hoelzl@49774: fixes f :: "'a \ ereal" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes H: "(\x. H (ereal (real (f x)))) \ borel_measurable M" hoelzl@49774: shows "(\x. H (f x)) \ borel_measurable M" hoelzl@49774: proof - hoelzl@50002: let ?F = "\x. if f x = \ then H \ else if f x = - \ then H (-\) else H (ereal (real (f x)))" hoelzl@49774: { fix x have "H (f x) = ?F x" by (cases "f x") auto } hoelzl@50002: with f H show ?thesis by simp hoelzl@47694: qed hoelzl@41981: hoelzl@49774: lemma hoelzl@50003: fixes f :: "'a \ ereal" assumes f[measurable]: "f \ borel_measurable M" hoelzl@50003: shows borel_measurable_ereal_abs[measurable(raw)]: "(\x. \f x\) \ borel_measurable M" hoelzl@50003: and borel_measurable_ereal_inverse[measurable(raw)]: "(\x. inverse (f x) :: ereal) \ borel_measurable M" hoelzl@50003: and borel_measurable_uminus_ereal[measurable(raw)]: "(\x. - f x :: ereal) \ borel_measurable M" hoelzl@49774: by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) hoelzl@49774: hoelzl@49774: lemma borel_measurable_uminus_eq_ereal[simp]: hoelzl@49774: "(\x. - f x :: ereal) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") hoelzl@49774: proof hoelzl@49774: assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp hoelzl@49774: qed auto hoelzl@49774: hoelzl@49774: lemma set_Collect_ereal2: hoelzl@49774: fixes f g :: "'a \ ereal" hoelzl@49774: assumes f: "f \ borel_measurable M" hoelzl@49774: assumes g: "g \ borel_measurable M" hoelzl@49774: assumes H: "{x \ space M. H (ereal (real (f x))) (ereal (real (g x)))} \ sets M" hoelzl@50002: "{x \ space borel. H (-\) (ereal x)} \ sets borel" hoelzl@50002: "{x \ space borel. H (\) (ereal x)} \ sets borel" hoelzl@50002: "{x \ space borel. H (ereal x) (-\)} \ sets borel" hoelzl@50002: "{x \ space borel. H (ereal x) (\)} \ sets borel" hoelzl@49774: shows "{x \ space M. H (f x) (g x)} \ sets M" hoelzl@49774: proof - hoelzl@50002: let ?G = "\y x. if g x = \ then H y \ else if g x = -\ then H y (-\) else H y (ereal (real (g x)))" hoelzl@50002: let ?F = "\x. if f x = \ then ?G \ x else if f x = -\ then ?G (-\) x else ?G (ereal (real (f x))) x" hoelzl@49774: { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } hoelzl@50002: note * = this hoelzl@50002: from assms show ?thesis hoelzl@50002: by (subst *) (simp del: space_borel split del: split_if) hoelzl@49774: qed hoelzl@49774: hoelzl@47694: lemma borel_measurable_ereal_iff: hoelzl@43920: shows "(\x. ereal (f x)) \ borel_measurable M \ f \ borel_measurable M" hoelzl@41981: proof hoelzl@43920: assume "(\x. ereal (f x)) \ borel_measurable M" hoelzl@43920: from borel_measurable_real_of_ereal[OF this] hoelzl@41981: show "f \ borel_measurable M" by auto hoelzl@41981: qed auto hoelzl@41981: hoelzl@47694: lemma borel_measurable_ereal_iff_real: hoelzl@43923: fixes f :: "'a \ ereal" hoelzl@43923: shows "f \ borel_measurable M \ hoelzl@41981: ((\x. real (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" hoelzl@41981: proof safe hoelzl@41981: assume *: "(\x. real (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" "f -` {-\} \ space M \ sets M" hoelzl@41981: have "f -` {\} \ space M = {x\space M. f x = \}" "f -` {-\} \ space M = {x\space M. f x = -\}" by auto hoelzl@41981: with * have **: "{x\space M. f x = \} \ sets M" "{x\space M. f x = -\} \ sets M" by simp_all wenzelm@46731: let ?f = "\x. if f x = \ then \ else if f x = -\ then -\ else ereal (real (f x))" hoelzl@41981: have "?f \ borel_measurable M" using * ** by (intro measurable_If) auto hoelzl@43920: also have "?f = f" by (auto simp: fun_eq_iff ereal_real) hoelzl@41981: finally show "f \ borel_measurable M" . hoelzl@50002: qed simp_all hoelzl@41830: hoelzl@47694: lemma borel_measurable_eq_atMost_ereal: hoelzl@43923: fixes f :: "'a \ ereal" hoelzl@43923: shows "f \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" hoelzl@41981: proof (intro iffI allI) hoelzl@41981: assume pos[rule_format]: "\a. f -` {..a} \ space M \ sets M" hoelzl@41981: show "f \ borel_measurable M" hoelzl@43920: unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le hoelzl@41981: proof (intro conjI allI) hoelzl@41981: fix a :: real hoelzl@43920: { fix x :: ereal assume *: "\i::nat. real i < x" hoelzl@41981: have "x = \" hoelzl@43920: proof (rule ereal_top) huffman@44666: fix B from reals_Archimedean2[of B] guess n .. hoelzl@43920: then have "ereal B < real n" by auto hoelzl@41981: with * show "B \ x" by (metis less_trans less_imp_le) hoelzl@41981: qed } hoelzl@41981: then have "f -` {\} \ space M = space M - (\i::nat. f -` {.. real i} \ space M)" hoelzl@41981: by (auto simp: not_le) hoelzl@50002: then show "f -` {\} \ space M \ sets M" using pos hoelzl@50002: by (auto simp del: UN_simps) hoelzl@41981: moreover hoelzl@43923: have "{-\::ereal} = {..-\}" by auto hoelzl@41981: then show "f -` {-\} \ space M \ sets M" using pos by auto hoelzl@43920: moreover have "{x\space M. f x \ ereal a} \ sets M" hoelzl@43920: using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute) hoelzl@41981: moreover have "{w \ space M. real (f w) \ a} = hoelzl@43920: (if a < 0 then {w \ space M. f w \ ereal a} - f -` {-\} \ space M hoelzl@43920: else {w \ space M. f w \ ereal a} \ (f -` {\} \ space M) \ (f -` {-\} \ space M))" (is "?l = ?r") hoelzl@41981: proof (intro set_eqI) fix x show "x \ ?l \ x \ ?r" by (cases "f x") auto qed hoelzl@41981: ultimately show "{w \ space M. real (f w) \ a} \ sets M" by auto hoelzl@35582: qed hoelzl@41981: qed (simp add: measurable_sets) hoelzl@35582: hoelzl@47694: lemma borel_measurable_eq_atLeast_ereal: hoelzl@43920: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" hoelzl@41981: proof hoelzl@41981: assume pos: "\a. f -` {a..} \ space M \ sets M" hoelzl@41981: moreover have "\a. (\x. - f x) -` {..a} = f -` {-a ..}" hoelzl@43920: by (auto simp: ereal_uminus_le_reorder) hoelzl@41981: ultimately have "(\x. - f x) \ borel_measurable M" hoelzl@43920: unfolding borel_measurable_eq_atMost_ereal by auto hoelzl@41981: then show "f \ borel_measurable M" by simp hoelzl@41981: qed (simp add: measurable_sets) hoelzl@35582: hoelzl@49774: lemma greater_eq_le_measurable: hoelzl@49774: fixes f :: "'a \ 'c::linorder" hoelzl@49774: shows "f -` {..< a} \ space M \ sets M \ f -` {a ..} \ space M \ sets M" hoelzl@49774: proof hoelzl@49774: assume "f -` {a ..} \ space M \ sets M" hoelzl@49774: moreover have "f -` {..< a} \ space M = space M - f -` {a ..} \ space M" by auto hoelzl@49774: ultimately show "f -` {..< a} \ space M \ sets M" by auto hoelzl@49774: next hoelzl@49774: assume "f -` {..< a} \ space M \ sets M" hoelzl@49774: moreover have "f -` {a ..} \ space M = space M - f -` {..< a} \ space M" by auto hoelzl@49774: ultimately show "f -` {a ..} \ space M \ sets M" by auto hoelzl@49774: qed hoelzl@49774: hoelzl@47694: lemma borel_measurable_ereal_iff_less: hoelzl@43920: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" hoelzl@43920: unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable .. hoelzl@38656: hoelzl@49774: lemma less_eq_ge_measurable: hoelzl@49774: fixes f :: "'a \ 'c::linorder" hoelzl@49774: shows "f -` {a <..} \ space M \ sets M \ f -` {..a} \ space M \ sets M" hoelzl@49774: proof hoelzl@49774: assume "f -` {a <..} \ space M \ sets M" hoelzl@49774: moreover have "f -` {..a} \ space M = space M - f -` {a <..} \ space M" by auto hoelzl@49774: ultimately show "f -` {..a} \ space M \ sets M" by auto hoelzl@49774: next hoelzl@49774: assume "f -` {..a} \ space M \ sets M" hoelzl@49774: moreover have "f -` {a <..} \ space M = space M - f -` {..a} \ space M" by auto hoelzl@49774: ultimately show "f -` {a <..} \ space M \ sets M" by auto hoelzl@49774: qed hoelzl@49774: hoelzl@47694: lemma borel_measurable_ereal_iff_ge: hoelzl@43920: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" hoelzl@43920: unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable .. hoelzl@38656: hoelzl@49774: lemma borel_measurable_ereal2: hoelzl@49774: fixes f g :: "'a \ ereal" hoelzl@41981: assumes f: "f \ borel_measurable M" hoelzl@41981: assumes g: "g \ borel_measurable M" hoelzl@49774: assumes H: "(\x. H (ereal (real (f x))) (ereal (real (g x)))) \ borel_measurable M" hoelzl@49774: "(\x. H (-\) (ereal (real (g x)))) \ borel_measurable M" hoelzl@49774: "(\x. H (\) (ereal (real (g x)))) \ borel_measurable M" hoelzl@49774: "(\x. H (ereal (real (f x))) (-\)) \ borel_measurable M" hoelzl@49774: "(\x. H (ereal (real (f x))) (\)) \ borel_measurable M" hoelzl@49774: shows "(\x. H (f x) (g x)) \ borel_measurable M" hoelzl@41981: proof - hoelzl@50002: let ?G = "\y x. if g x = \ then H y \ else if g x = - \ then H y (-\) else H y (ereal (real (g x)))" hoelzl@50002: let ?F = "\x. if f x = \ then ?G \ x else if f x = - \ then ?G (-\) x else ?G (ereal (real (f x))) x" hoelzl@49774: { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } hoelzl@50002: note * = this hoelzl@50002: from assms show ?thesis unfolding * by simp hoelzl@41981: qed hoelzl@41981: hoelzl@49774: lemma hoelzl@49774: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" hoelzl@49774: shows borel_measurable_ereal_eq_const: "{x\space M. f x = c} \ sets M" hoelzl@49774: and borel_measurable_ereal_neq_const: "{x\space M. f x \ c} \ sets M" hoelzl@49774: using f by auto hoelzl@38656: hoelzl@50003: lemma [measurable(raw)]: hoelzl@43920: fixes f :: "'a \ ereal" hoelzl@50003: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" hoelzl@50002: shows borel_measurable_ereal_add: "(\x. f x + g x) \ borel_measurable M" hoelzl@50002: and borel_measurable_ereal_times: "(\x. f x * g x) \ borel_measurable M" hoelzl@50002: and borel_measurable_ereal_min: "(\x. min (g x) (f x)) \ borel_measurable M" hoelzl@50002: and borel_measurable_ereal_max: "(\x. max (g x) (f x)) \ borel_measurable M" hoelzl@50003: by (simp_all add: borel_measurable_ereal2 min_def max_def) hoelzl@49774: hoelzl@50003: lemma [measurable(raw)]: hoelzl@49774: fixes f g :: "'a \ ereal" hoelzl@49774: assumes "f \ borel_measurable M" hoelzl@49774: assumes "g \ borel_measurable M" hoelzl@50002: shows borel_measurable_ereal_diff: "(\x. f x - g x) \ borel_measurable M" hoelzl@50002: and borel_measurable_ereal_divide: "(\x. f x / g x) \ borel_measurable M" hoelzl@50003: using assms by (simp_all add: minus_ereal_def divide_ereal_def) hoelzl@38656: hoelzl@50003: lemma borel_measurable_ereal_setsum[measurable (raw)]: hoelzl@43920: fixes f :: "'c \ 'a \ ereal" hoelzl@41096: assumes "\i. i \ S \ f i \ borel_measurable M" hoelzl@41096: shows "(\x. \i\S. f i x) \ borel_measurable M" hoelzl@41096: proof cases hoelzl@41096: assume "finite S" hoelzl@41096: thus ?thesis using assms hoelzl@41096: by induct auto hoelzl@49774: qed simp hoelzl@38656: hoelzl@50003: lemma borel_measurable_ereal_setprod[measurable (raw)]: hoelzl@43920: fixes f :: "'c \ 'a \ ereal" hoelzl@38656: assumes "\i. i \ S \ f i \ borel_measurable M" hoelzl@41096: shows "(\x. \i\S. f i x) \ borel_measurable M" hoelzl@38656: proof cases hoelzl@38656: assume "finite S" hoelzl@41096: thus ?thesis using assms by induct auto hoelzl@41096: qed simp hoelzl@38656: hoelzl@50003: lemma borel_measurable_SUP[measurable (raw)]: hoelzl@43920: fixes f :: "'d\countable \ 'a \ ereal" hoelzl@38656: assumes "\i. i \ A \ f i \ borel_measurable M" hoelzl@41097: shows "(\x. SUP i : A. f i x) \ borel_measurable M" (is "?sup \ borel_measurable M") hoelzl@43920: unfolding borel_measurable_ereal_iff_ge hoelzl@41981: proof hoelzl@38656: fix a hoelzl@41981: have "?sup -` {a<..} \ space M = (\i\A. {x\space M. a < f i x})" noschinl@46884: by (auto simp: less_SUP_iff) hoelzl@41981: then show "?sup -` {a<..} \ space M \ sets M" hoelzl@38656: using assms by auto hoelzl@38656: qed hoelzl@38656: hoelzl@50003: lemma borel_measurable_INF[measurable (raw)]: hoelzl@43920: fixes f :: "'d :: countable \ 'a \ ereal" hoelzl@38656: assumes "\i. i \ A \ f i \ borel_measurable M" hoelzl@41097: shows "(\x. INF i : A. f i x) \ borel_measurable M" (is "?inf \ borel_measurable M") hoelzl@43920: unfolding borel_measurable_ereal_iff_less hoelzl@41981: proof hoelzl@38656: fix a hoelzl@41981: have "?inf -` {.. space M = (\i\A. {x\space M. f i x < a})" noschinl@46884: by (auto simp: INF_less_iff) hoelzl@41981: then show "?inf -` {.. space M \ sets M" hoelzl@38656: using assms by auto hoelzl@38656: qed hoelzl@38656: hoelzl@50003: lemma [measurable (raw)]: hoelzl@43920: fixes f :: "nat \ 'a \ ereal" hoelzl@41981: assumes "\i. f i \ borel_measurable M" hoelzl@50002: shows borel_measurable_liminf: "(\x. liminf (\i. f i x)) \ borel_measurable M" hoelzl@50002: and borel_measurable_limsup: "(\x. limsup (\i. f i x)) \ borel_measurable M" haftmann@56212: unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto hoelzl@35692: hoelzl@50104: lemma sets_Collect_eventually_sequentially[measurable]: hoelzl@50003: "(\i. {x\space M. P x i} \ sets M) \ {x\space M. eventually (P x) sequentially} \ sets M" hoelzl@50003: unfolding eventually_sequentially by simp hoelzl@50003: hoelzl@50003: lemma sets_Collect_ereal_convergent[measurable]: hoelzl@50003: fixes f :: "nat \ 'a => ereal" hoelzl@50003: assumes f[measurable]: "\i. f i \ borel_measurable M" hoelzl@50003: shows "{x\space M. convergent (\i. f i x)} \ sets M" hoelzl@50003: unfolding convergent_ereal by auto hoelzl@50003: hoelzl@50003: lemma borel_measurable_extreal_lim[measurable (raw)]: hoelzl@50003: fixes f :: "nat \ 'a \ ereal" hoelzl@50003: assumes [measurable]: "\i. f i \ borel_measurable M" hoelzl@50003: shows "(\x. lim (\i. f i x)) \ borel_measurable M" hoelzl@50003: proof - hoelzl@50003: have "\x. lim (\i. f i x) = (if convergent (\i. f i x) then limsup (\i. f i x) else (THE i. False))" hoelzl@51351: by (simp add: lim_def convergent_def convergent_limsup_cl) hoelzl@50003: then show ?thesis hoelzl@50003: by simp hoelzl@50003: qed hoelzl@50003: hoelzl@49774: lemma borel_measurable_ereal_LIMSEQ: hoelzl@49774: fixes u :: "nat \ 'a \ ereal" hoelzl@49774: assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" hoelzl@49774: and u: "\i. u i \ borel_measurable M" hoelzl@49774: shows "u' \ borel_measurable M" hoelzl@47694: proof - hoelzl@49774: have "\x. x \ space M \ u' x = liminf (\n. u n x)" hoelzl@49774: using u' by (simp add: lim_imp_Liminf[symmetric]) hoelzl@50003: with u show ?thesis by (simp cong: measurable_cong) hoelzl@47694: qed hoelzl@47694: hoelzl@50003: lemma borel_measurable_extreal_suminf[measurable (raw)]: hoelzl@43920: fixes f :: "nat \ 'a \ ereal" hoelzl@50003: assumes [measurable]: "\i. f i \ borel_measurable M" hoelzl@41981: shows "(\x. (\i. f i x)) \ borel_measurable M" hoelzl@50003: unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp hoelzl@39092: hoelzl@56994: subsection {* LIMSEQ is borel measurable *} hoelzl@39092: hoelzl@47694: lemma borel_measurable_LIMSEQ: hoelzl@39092: fixes u :: "nat \ 'a \ real" hoelzl@39092: assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" hoelzl@39092: and u: "\i. u i \ borel_measurable M" hoelzl@39092: shows "u' \ borel_measurable M" hoelzl@39092: proof - hoelzl@43920: have "\x. x \ space M \ liminf (\n. ereal (u n x)) = ereal (u' x)" wenzelm@46731: using u' by (simp add: lim_imp_Liminf) hoelzl@43920: moreover from u have "(\x. liminf (\n. ereal (u n x))) \ borel_measurable M" hoelzl@39092: by auto hoelzl@43920: ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) hoelzl@39092: qed hoelzl@39092: hoelzl@56993: lemma borel_measurable_LIMSEQ_metric: hoelzl@56993: fixes f :: "nat \ 'a \ 'b :: metric_space" hoelzl@56993: assumes [measurable]: "\i. f i \ borel_measurable M" hoelzl@56993: assumes lim: "\x. x \ space M \ (\i. f i x) ----> g x" hoelzl@56993: shows "g \ borel_measurable M" hoelzl@56993: unfolding borel_eq_closed hoelzl@56993: proof (safe intro!: measurable_measure_of) hoelzl@56993: fix A :: "'b set" assume "closed A" hoelzl@56993: hoelzl@56993: have [measurable]: "(\x. infdist (g x) A) \ borel_measurable M" hoelzl@56993: proof (rule borel_measurable_LIMSEQ) hoelzl@56993: show "\x. x \ space M \ (\i. infdist (f i x) A) ----> infdist (g x) A" hoelzl@56993: by (intro tendsto_infdist lim) hoelzl@56993: show "\i. (\x. infdist (f i x) A) \ borel_measurable M" hoelzl@56993: by (intro borel_measurable_continuous_on[where f="\x. infdist x A"] hoelzl@56993: continuous_at_imp_continuous_on ballI continuous_infdist isCont_ident) auto hoelzl@56993: qed hoelzl@56993: hoelzl@56993: show "g -` A \ space M \ sets M" hoelzl@56993: proof cases hoelzl@56993: assume "A \ {}" hoelzl@56993: then have "\x. infdist x A = 0 \ x \ A" hoelzl@56993: using `closed A` by (simp add: in_closed_iff_infdist_zero) hoelzl@56993: then have "g -` A \ space M = {x\space M. infdist (g x) A = 0}" hoelzl@56993: by auto hoelzl@56993: also have "\ \ sets M" hoelzl@56993: by measurable hoelzl@56993: finally show ?thesis . hoelzl@56993: qed simp hoelzl@56993: qed auto hoelzl@56993: hoelzl@50002: lemma sets_Collect_Cauchy[measurable]: hoelzl@57036: fixes f :: "nat \ 'a => 'b::{metric_space, second_countable_topology}" hoelzl@50002: assumes f[measurable]: "\i. f i \ borel_measurable M" hoelzl@49774: shows "{x\space M. Cauchy (\i. f i x)} \ sets M" hoelzl@57036: unfolding metric_Cauchy_iff2 using f by auto hoelzl@49774: hoelzl@50002: lemma borel_measurable_lim[measurable (raw)]: hoelzl@57036: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" hoelzl@50002: assumes f[measurable]: "\i. f i \ borel_measurable M" hoelzl@49774: shows "(\x. lim (\i. f i x)) \ borel_measurable M" hoelzl@49774: proof - hoelzl@50002: def u' \ "\x. lim (\i. if Cauchy (\i. f i x) then f i x else 0)" hoelzl@50002: then have *: "\x. lim (\i. f i x) = (if Cauchy (\i. f i x) then u' x else (THE x. False))" hoelzl@49774: by (auto simp: lim_def convergent_eq_cauchy[symmetric]) hoelzl@50002: have "u' \ borel_measurable M" hoelzl@57036: proof (rule borel_measurable_LIMSEQ_metric) hoelzl@50002: fix x hoelzl@50002: have "convergent (\i. if Cauchy (\i. f i x) then f i x else 0)" hoelzl@49774: by (cases "Cauchy (\i. f i x)") hoelzl@50002: (auto simp add: convergent_eq_cauchy[symmetric] convergent_def) hoelzl@50002: then show "(\i. if Cauchy (\i. f i x) then f i x else 0) ----> u' x" hoelzl@50002: unfolding u'_def hoelzl@50002: by (rule convergent_LIMSEQ_iff[THEN iffD1]) hoelzl@50002: qed measurable hoelzl@50002: then show ?thesis hoelzl@50002: unfolding * by measurable hoelzl@49774: qed hoelzl@49774: hoelzl@50002: lemma borel_measurable_suminf[measurable (raw)]: hoelzl@57036: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" hoelzl@50002: assumes f[measurable]: "\i. f i \ borel_measurable M" hoelzl@49774: shows "(\x. suminf (\i. f i x)) \ borel_measurable M" hoelzl@50002: unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp hoelzl@49774: hoelzl@57447: (* Proof by Jeremy Avigad and Luke Serafin *) hoelzl@57447: lemma isCont_borel: hoelzl@57447: fixes f :: "'b::metric_space \ 'a::metric_space" hoelzl@57447: shows "{x. isCont f x} \ sets borel" hoelzl@57447: proof - hoelzl@57447: let ?I = "\j. inverse(real (Suc j))" hoelzl@57447: hoelzl@57447: { fix x hoelzl@57447: have "isCont f x = (\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i)" hoelzl@57447: unfolding continuous_at_eps_delta hoelzl@57447: proof safe hoelzl@57447: fix i assume "\e>0. \d>0. \y. dist y x < d \ dist (f y) (f x) < e" hoelzl@57447: moreover have "0 < ?I i / 2" hoelzl@57447: by simp hoelzl@57447: ultimately obtain d where d: "0 < d" "\y. dist x y < d \ dist (f y) (f x) < ?I i / 2" hoelzl@57447: by (metis dist_commute) hoelzl@57447: then obtain j where j: "?I j < d" hoelzl@57447: by (metis reals_Archimedean) hoelzl@57447: hoelzl@57447: show "\j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" hoelzl@57447: proof (safe intro!: exI[where x=j]) hoelzl@57447: fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" hoelzl@57447: have "dist (f y) (f z) \ dist (f y) (f x) + dist (f z) (f x)" hoelzl@57447: by (rule dist_triangle2) hoelzl@57447: also have "\ < ?I i / 2 + ?I i / 2" hoelzl@57447: by (intro add_strict_mono d less_trans[OF _ j] *) hoelzl@57447: also have "\ \ ?I i" hoelzl@57447: by (simp add: field_simps real_of_nat_Suc) hoelzl@57447: finally show "dist (f y) (f z) \ ?I i" hoelzl@57447: by simp hoelzl@57447: qed hoelzl@57447: next hoelzl@57447: fix e::real assume "0 < e" hoelzl@57447: then obtain n where n: "?I n < e" hoelzl@57447: by (metis reals_Archimedean) hoelzl@57447: assume "\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" hoelzl@57447: from this[THEN spec, of "Suc n"] hoelzl@57447: obtain j where j: "\y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I (Suc n)" hoelzl@57447: by auto hoelzl@57447: hoelzl@57447: show "\d>0. \y. dist y x < d \ dist (f y) (f x) < e" hoelzl@57447: proof (safe intro!: exI[of _ "?I j"]) hoelzl@57447: fix y assume "dist y x < ?I j" hoelzl@57447: then have "dist (f y) (f x) \ ?I (Suc n)" hoelzl@57447: by (intro j) (auto simp: dist_commute) hoelzl@57447: also have "?I (Suc n) < ?I n" hoelzl@57447: by simp hoelzl@57447: also note n hoelzl@57447: finally show "dist (f y) (f x) < e" . hoelzl@57447: qed simp hoelzl@57447: qed } hoelzl@57447: note * = this hoelzl@57447: hoelzl@57447: have **: "\e y. open {x. dist x y < e}" hoelzl@57447: using open_ball by (simp_all add: ball_def dist_commute) hoelzl@57447: hoelzl@57447: have "{x\space borel. isCont f x } \ sets borel" hoelzl@57447: unfolding * hoelzl@57447: apply (intro sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex) hoelzl@57447: apply (simp add: Collect_all_eq) hoelzl@57447: apply (intro borel_closed closed_INT ballI closed_Collect_imp open_Collect_conj **) hoelzl@57447: apply auto hoelzl@57447: done hoelzl@57447: then show ?thesis hoelzl@57447: by simp hoelzl@57447: qed hoelzl@57447: immler@54775: no_notation immler@54775: eucl_less (infix "