src/HOL/Probability/Borel.thy
changeset 38656 d5d342611edb
parent 37887 2ae085b07f2f
child 38705 aaee86c0e237
--- a/src/HOL/Probability/Borel.thy	Mon Aug 23 17:46:13 2010 +0200
+++ b/src/HOL/Probability/Borel.thy	Mon Aug 23 19:35:57 2010 +0200
@@ -1,242 +1,199 @@
-header {*Borel Sets*}
+(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
+
+header {*Borel spaces*}
 
 theory Borel
-  imports Measure
+  imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
 begin
 
-text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
-
-definition borel_space where
-  "borel_space = sigma (UNIV::real set) (range (\<lambda>a::real. {x. x \<le> a}))"
+section "Generic Borel spaces"
 
-definition borel_measurable where
-  "borel_measurable a = measurable a borel_space"
+definition "borel_space = sigma (UNIV::'a::topological_space set) open"
+abbreviation "borel_measurable M \<equiv> measurable M borel_space"
 
-definition indicator_fn where
-  "indicator_fn s = (\<lambda>x. if x \<in> s then 1 else (0::real))"
+interpretation borel_space: sigma_algebra borel_space
+  using sigma_algebra_sigma by (auto simp: borel_space_def)
 
 lemma in_borel_measurable:
    "f \<in> borel_measurable M \<longleftrightarrow>
-    sigma_algebra M \<and>
-    (\<forall>s \<in> sets (sigma UNIV (range (\<lambda>a::real. {x. x \<le> a}))).
-      f -` s \<inter> space M \<in> sets M)"
-apply (auto simp add: borel_measurable_def measurable_def borel_space_def) 
-apply (metis PowD UNIV_I Un_commute sigma_algebra_sigma subset_Pow_Union subset_UNIV subset_Un_eq) 
-done
-
-lemma (in sigma_algebra) borel_measurable_const:
-   "(\<lambda>x. c) \<in> borel_measurable M"
-  by (auto simp add: in_borel_measurable prems)
-
-lemma (in sigma_algebra) borel_measurable_indicator:
-  assumes a: "a \<in> sets M"
-  shows "indicator_fn a \<in> borel_measurable M"
-apply (auto simp add: in_borel_measurable indicator_fn_def prems)
-apply (metis Diff_eq Int_commute a compl_sets) 
-done
+    (\<forall>S \<in> sets (sigma UNIV open).
+      f -` S \<inter> space M \<in> sets M)"
+  by (auto simp add: measurable_def borel_space_def)
 
-lemma Collect_eq: "{w \<in> X. f w \<le> a} = {w. f w \<le> a} \<inter> X"
-  by (metis Collect_conj_eq Collect_mem_eq Int_commute)
+lemma in_borel_measurable_borel_space:
+   "f \<in> borel_measurable M \<longleftrightarrow>
+    (\<forall>S \<in> sets borel_space.
+      f -` S \<inter> space M \<in> sets M)"
+  by (auto simp add: measurable_def borel_space_def)
 
-lemma (in measure_space) borel_measurable_le_iff:
-   "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
-proof 
-  assume f: "f \<in> borel_measurable M"
-  { fix a
-    have "{w \<in> space M. f w \<le> a} \<in> sets M" using f
-      apply (auto simp add: in_borel_measurable sigma_def Collect_eq)
-      apply (drule_tac x="{x. x \<le> a}" in bspec, auto)
-      apply (metis equalityE rangeI subsetD sigma_sets.Basic)  
-      done
-    }
-  thus "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M" by blast
-next
-  assume "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M"
-  thus "f \<in> borel_measurable M" 
-    apply (simp add: borel_measurable_def borel_space_def Collect_eq) 
-    apply (rule measurable_sigma, auto) 
-    done
+lemma space_borel_space[simp]: "space borel_space = UNIV"
+  unfolding borel_space_def by auto
+
+lemma borel_space_open[simp]:
+  assumes "open A" shows "A \<in> sets borel_space"
+proof -
+  have "A \<in> open" unfolding mem_def using assms .
+  thus ?thesis unfolding borel_space_def sigma_def by (auto intro!: sigma_sets.Basic)
 qed
 
-lemma Collect_less_le:
-     "{w \<in> X. f w < g w} = (\<Union>n. {w \<in> X. f w \<le> g w - inverse(real(Suc n))})"
-  proof auto
-    fix w
-    assume w: "f w < g w"
-    hence nz: "g w - f w \<noteq> 0"
-      by arith
-    with w have "real(Suc(natceiling(inverse(g w - f w)))) > inverse(g w - f w)"
-      by (metis lessI order_le_less_trans real_natceiling_ge real_of_nat_less_iff)       hence "inverse(real(Suc(natceiling(inverse(g w - f w)))))
-             < inverse(inverse(g w - f w))" 
-      by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_linordered_idom nz positive_imp_inverse_positive order_antisym less_le w)
-    hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < g w - f w"
-      by (metis inverse_inverse_eq order_less_le_trans order_refl)
-    thus "\<exists>n. f w \<le> g w - inverse(real(Suc n))" using w
-      by (rule_tac x="natceiling(inverse(g w - f w))" in exI, auto)
-  next
-    fix w n
-    assume le: "f w \<le> g w - inverse(real(Suc n))"
-    hence "0 < inverse(real(Suc n))"
-      by simp
-    thus "f w < g w" using le
-      by arith 
-  qed
-
-
-lemma (in sigma_algebra) sigma_le_less:
-  assumes M: "!!a::real. {w \<in> space M. f w \<le> a} \<in> sets M"
-  shows "{w \<in> space M. f w < a} \<in> sets M"
+lemma borel_space_closed[simp]:
+  assumes "closed A" shows "A \<in> sets borel_space"
 proof -
-  show ?thesis using Collect_less_le [of "space M" f "\<lambda>x. a"]
-    by (auto simp add: countable_UN M) 
+  have "space borel_space - (- A) \<in> sets borel_space"
+    using assms unfolding closed_def by (blast intro: borel_space_open)
+  thus ?thesis by simp
 qed
 
-lemma (in sigma_algebra) sigma_less_ge:
-  assumes M: "!!a::real. {w \<in> space M. f w < a} \<in> sets M"
-  shows "{w \<in> space M. a \<le> f w} \<in> sets M"
-proof -
-  have "{w \<in> space M. a \<le> f w} = space M - {w \<in> space M. f w < a}"
-    by auto
-  thus ?thesis using M
-    by auto
-qed
-
-lemma (in sigma_algebra) sigma_ge_gr:
-  assumes M: "!!a::real. {w \<in> space M. a \<le> f w} \<in> sets M"
-  shows "{w \<in> space M. a < f w} \<in> sets M"
-proof -
-  show ?thesis using Collect_less_le [of "space M" "\<lambda>x. a" f]
-    by (auto simp add: countable_UN le_diff_eq M) 
+lemma (in sigma_algebra) borel_measurable_vimage:
+  fixes f :: "'a \<Rightarrow> 'x::t2_space"
+  assumes borel: "f \<in> borel_measurable M"
+  shows "f -` {x} \<inter> space M \<in> sets M"
+proof (cases "x \<in> f ` space M")
+  case True then obtain y where "x = f y" by auto
+  from closed_sing[of "f y"]
+  have "{f y} \<in> sets borel_space" by (rule borel_space_closed)
+  with assms show ?thesis
+    unfolding in_borel_measurable_borel_space `x = f y` by auto
+next
+  case False hence "f -` {x} \<inter> space M = {}" by auto
+  thus ?thesis by auto
 qed
 
-lemma (in sigma_algebra) sigma_gr_le:
-  assumes M: "!!a::real. {w \<in> space M. a < f w} \<in> sets M"
-  shows "{w \<in> space M. f w \<le> a} \<in> sets M"
-proof -
-  have "{w \<in> space M. f w \<le> a} = space M - {w \<in> space M. a < f w}" 
-    by auto
-  thus ?thesis
-    by (simp add: M compl_sets)
-qed
+lemma (in sigma_algebra) borel_measurableI:
+  fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
+  assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+  shows "f \<in> borel_measurable M"
+  unfolding borel_space_def
+proof (rule measurable_sigma)
+  fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M"
+    using assms[of S] by (simp add: mem_def)
+qed simp_all
 
-lemma (in measure_space) borel_measurable_gr_iff:
-   "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
-proof (auto simp add: borel_measurable_le_iff sigma_gr_le) 
-  fix u
-  assume M: "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M"
-  have "{w \<in> space M. u < f w} = space M - {w \<in> space M. f w \<le> u}"
-    by auto
-  thus "{w \<in> space M. u < f w} \<in> sets M"
-    by (force simp add: compl_sets countable_UN M)
-qed
+lemma borel_space_singleton[simp, intro]:
+  fixes x :: "'a::t1_space"
+  shows "A \<in> sets borel_space \<Longrightarrow> insert x A \<in> sets borel_space"
+  proof (rule borel_space.insert_in_sets)
+    show "{x} \<in> sets borel_space"
+      using closed_sing[of x] by (rule borel_space_closed)
+  qed simp
+
+lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
+  "(\<lambda>x. c) \<in> borel_measurable M"
+  by (auto intro!: measurable_const)
 
-lemma (in measure_space) borel_measurable_less_iff:
-   "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
-proof (auto simp add: borel_measurable_le_iff sigma_le_less) 
-  fix u
-  assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M"
-  have "{w \<in> space M. f w \<le> u} = space M - {w \<in> space M. u < f w}"
-    by auto
-  thus "{w \<in> space M. f w \<le> u} \<in> sets M"
-    using Collect_less_le [of "space M" "\<lambda>x. u" f] 
-    by (force simp add: compl_sets countable_UN le_diff_eq sigma_less_ge M)
-qed
+lemma (in sigma_algebra) borel_measurable_indicator:
+  assumes A: "A \<in> sets M"
+  shows "indicator A \<in> borel_measurable M"
+  unfolding indicator_def_raw using A
+  by (auto intro!: measurable_If_set borel_measurable_const)
 
-lemma (in measure_space) borel_measurable_ge_iff:
-   "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
-proof (auto simp add: borel_measurable_less_iff sigma_le_less sigma_ge_gr sigma_gr_le) 
-  fix u
-  assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M"
-  have "{w \<in> space M. u \<le> f w} = space M - {w \<in> space M. f w < u}"
-    by auto
-  thus "{w \<in> space M. u \<le> f w} \<in> sets M"
-    by (force simp add: compl_sets countable_UN M)
+lemma borel_measurable_translate:
+  assumes "A \<in> sets borel_space" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel_space"
+  shows "f -` A \<in> sets borel_space"
+proof -
+  have "A \<in> sigma_sets UNIV open" using assms
+    by (simp add: borel_space_def sigma_def)
+  thus ?thesis
+  proof (induct rule: sigma_sets.induct)
+    case (Basic a) thus ?case using trans[of a] by (simp add: mem_def)
+  next
+    case (Compl a)
+    moreover have "UNIV \<in> sets borel_space"
+      by (metis borel_space.top borel_space_def_raw mem_def space_sigma)
+    ultimately show ?case
+      by (auto simp: vimage_Diff borel_space.Diff)
+  qed (auto simp add: vimage_UN)
 qed
 
-lemma (in measure_space) affine_borel_measurable:
-  assumes g: "g \<in> borel_measurable M"
-  shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
-proof (cases rule: linorder_cases [of b 0])
-  case equal thus ?thesis
-    by (simp add: borel_measurable_const)
-next
-  case less
-    {
-      fix w c
-      have "a + g w * b \<le> c \<longleftrightarrow> g w * b \<le> c - a"
-        by auto
-      also have "... \<longleftrightarrow> (c-a)/b \<le> g w" using less
-        by (metis divide_le_eq less less_asym)
-      finally have "a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" .
-    }
-    hence "\<And>w c. a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" .
-    thus ?thesis using less g
-      by (simp add: borel_measurable_ge_iff [of g]) 
-         (simp add: borel_measurable_le_iff)
-next
-  case greater
-    hence 0: "\<And>x c. (g x * b \<le> c - a) \<longleftrightarrow> (g x \<le> (c - a) / b)"
-      by (metis mult_imp_le_div_pos le_divide_eq) 
-    have 1: "\<And>x c. (a + g x * b \<le> c) \<longleftrightarrow> (g x * b \<le> c - a)"
-      by auto
-    from greater g
-    show ?thesis
-      by (simp add: borel_measurable_le_iff 0 1) 
-qed
+section "Borel spaces on euclidean spaces"
+
+lemma lessThan_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{..< a} \<in> sets borel_space"
+  by (blast intro: borel_space_open)
+
+lemma greaterThan_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a <..} \<in> sets borel_space"
+  by (blast intro: borel_space_open)
+
+lemma greaterThanLessThan_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a<..<b} \<in> sets borel_space"
+  by (blast intro: borel_space_open)
+
+lemma atMost_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{..a} \<in> sets borel_space"
+  by (blast intro: borel_space_closed)
+
+lemma atLeast_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..} \<in> sets borel_space"
+  by (blast intro: borel_space_closed)
+
+lemma atLeastAtMost_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..b} \<in> sets borel_space"
+  by (blast intro: borel_space_closed)
 
-definition
-  nat_to_rat_surj :: "nat \<Rightarrow> rat" where
- "nat_to_rat_surj n = (let (i,j) = prod_decode n
-                       in Fract (int_decode i) (int_decode j))"
+lemma greaterThanAtMost_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a<..b} \<in> sets borel_space"
+  unfolding greaterThanAtMost_def by blast
+
+lemma atLeastLessThan_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..<b} \<in> sets borel_space"
+  unfolding atLeastLessThan_def by blast
+
+lemma hafspace_less_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel_space"
+  by (auto intro!: borel_space_open open_halfspace_component_gt)
 
-lemma nat_to_rat_surj: "surj nat_to_rat_surj"
-proof (auto simp add: surj_def nat_to_rat_surj_def) 
-  fix y
-  show "\<exists>x. y = (\<lambda>(i, j). Fract (int_decode i) (int_decode j)) (prod_decode x)"
-  proof (cases y)
-    case (Fract a b)
-      obtain i where i: "int_decode i = a" using surj_int_decode
-        by (metis surj_def) 
-      obtain j where j: "int_decode j = b" using surj_int_decode
-        by (metis surj_def)
-      obtain n where n: "prod_decode n = (i,j)" using surj_prod_decode
-        by (metis surj_def)
+lemma hafspace_greater_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel_space"
+  by (auto intro!: borel_space_open open_halfspace_component_lt)
 
-      from Fract i j n show ?thesis
-        by (metis prod.cases)
-  qed
-qed
+lemma hafspace_less_eq_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel_space"
+  by (auto intro!: borel_space_closed closed_halfspace_component_ge)
 
-lemma rats_enumeration: "\<rat> = range (of_rat o nat_to_rat_surj)" 
-  using nat_to_rat_surj
-  by (auto simp add: image_def surj_def) (metis Rats_cases) 
+lemma hafspace_greater_eq_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel_space"
+  by (auto intro!: borel_space_closed closed_halfspace_component_le)
 
-lemma (in measure_space) borel_measurable_less_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_less[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "{w \<in> space M. f w < g w} \<in> sets M"
 proof -
   have "{w \<in> space M. f w < g w} =
-        (\<Union>r\<in>\<rat>. {w \<in> space M. f w < r} \<inter> {w \<in> space M. r < g w })"
-    by (auto simp add: Rats_dense_in_real)
-  thus ?thesis using f g 
-    by (simp add: borel_measurable_less_iff [of f]  
-                  borel_measurable_gr_iff [of g]) 
-       (blast intro: gen_countable_UN [OF rats_enumeration])
+        (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
+    using Rats_dense_in_real by (auto simp add: Rats_def)
+  then show ?thesis using f g
+    by simp (blast intro: measurable_sets)
 qed
- 
-lemma (in measure_space) borel_measurable_leq_borel_measurable:
+
+lemma (in sigma_algebra) borel_measurable_le[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
 proof -
-  have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}" 
-    by auto 
-  thus ?thesis using f g 
-    by (simp add: borel_measurable_less_borel_measurable compl_sets)
+  have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
+    by auto
+  thus ?thesis using f g
+    by simp blast
 qed
 
-lemma (in measure_space) borel_measurable_eq_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_eq[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "{w \<in> space M. f w = g w} \<in> sets M"
@@ -244,40 +201,512 @@
   have "{w \<in> space M. f w = g w} =
         {w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
     by auto
-  thus ?thesis using f g 
-    by (simp add: borel_measurable_leq_borel_measurable Int) 
+  thus ?thesis using f g by auto
 qed
 
-lemma (in measure_space) borel_measurable_neq_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_neq[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 proof -
   have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}"
     by auto
-  thus ?thesis using f g 
-    by (simp add: borel_measurable_eq_borel_measurable compl_sets) 
+  thus ?thesis using f g by auto
+qed
+
+subsection "Borel space equals sigma algebras over intervals"
+
+lemma rational_boxes:
+  fixes x :: "'a\<Colon>ordered_euclidean_space"
+  assumes "0 < e"
+  shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
+proof -
+  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
+  then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
+  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
+  proof
+    fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
+    show "?th i" by auto
+  qed
+  from choice[OF this] guess a .. note a = this
+  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
+  proof
+    fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
+    show "?th i" by auto
+  qed
+  from choice[OF this] guess b .. note b = this
+  { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
+    have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
+      unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
+    also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
+    proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
+      fix i assume i: "i \<in> {..<DIM('a)}"
+      have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
+      moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
+      moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
+      ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
+      then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
+        unfolding e'_def by (auto simp: dist_real_def)
+      then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
+        by (rule power_strict_mono) auto
+      then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
+        by (simp add: power_divide)
+    qed auto
+    also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
+    finally have "dist x y < e" . }
+  with a b show ?thesis
+    apply (rule_tac exI[of _ "Chi a"])
+    apply (rule_tac exI[of _ "Chi b"])
+    using eucl_less[where 'a='a] by auto
+qed
+
+lemma ex_rat_list:
+  fixes x :: "'a\<Colon>ordered_euclidean_space"
+  assumes "\<And> i. x $$ i \<in> \<rat>"
+  shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
+proof -
+  have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
+  from choice[OF this] guess r ..
+  then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
+qed
+
+lemma open_UNION:
+  fixes M :: "'a\<Colon>ordered_euclidean_space set"
+  assumes "open M"
+  shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
+                   (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
+    (is "M = UNION ?idx ?box")
+proof safe
+  fix x assume "x \<in> M"
+  obtain e where e: "e > 0" "ball x e \<subseteq> M"
+    using openE[OF assms `x \<in> M`] by auto
+  then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
+    using rational_boxes[OF e(1)] by blast
+  then obtain p q where pq: "length p = DIM ('a)"
+                            "length q = DIM ('a)"
+                            "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
+    using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
+  hence p: "Chi (of_rat \<circ> op ! p) = a"
+    using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
+    unfolding o_def by auto
+  from pq have q: "Chi (of_rat \<circ> op ! q) = b"
+    using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
+    unfolding o_def by auto
+  have "x \<in> ?box (p, q)"
+    using p q ab by auto
+  thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
+qed auto
+
+lemma halfspace_span_open:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))
+    \<subseteq> sets borel_space"
+  by (auto intro!: borel_space.sigma_sets_subset[simplified] borel_space_open
+                   open_halfspace_component_lt simp: sets_sigma)
+
+lemma halfspace_lt_in_halfspace:
+  "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
+  unfolding sets_sigma by (rule sigma_sets.Basic) auto
+
+lemma halfspace_gt_in_halfspace:
+  "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
+    (is "?set \<in> sets ?SIGMA")
+proof -
+  interpret sigma_algebra ?SIGMA by (rule sigma_algebra_sigma) simp
+  have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
+  proof (safe, simp_all add: not_less)
+    fix x 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)
+    then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
+      by (blast intro: less_imp_le)
+  next
+    fix x n
+    have "a < a + 1 / real (Suc n)" by auto
+    also assume "\<dots> \<le> x"
+    finally show "a < x" .
+  qed
+  show "?set \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace)
 qed
 
-lemma (in measure_space) borel_measurable_add_borel_measurable:
+lemma (in sigma_algebra) sets_sigma_subset:
+  assumes "A = space M"
+  assumes "B \<subseteq> sets M"
+  shows "sets (sigma A B) \<subseteq> sets M"
+  by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms)
+
+lemma open_span_halfspace:
+  "sets borel_space \<subseteq> sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})))"
+    (is "_ \<subseteq> sets ?SIGMA")
+proof (unfold borel_space_def, rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp
+  then interpret sigma_algebra ?SIGMA .
+  fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def .
+  from open_UNION[OF this]
+  obtain I where *: "S =
+    (\<Union>(a, b)\<in>I.
+        (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
+        (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
+    unfolding greaterThanLessThan_def
+    unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
+    unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
+    by blast
+  show "S \<in> sets ?SIGMA"
+    unfolding *
+    by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace)
+qed auto
+
+lemma halfspace_span_halfspace_le:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq>
+   sets (sigma UNIV (range (\<lambda> (a, i). {x. x $$ i \<le> a})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a i
+  have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
+  proof (safe, simp_all)
+    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)
+    then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
+      by (blast intro: less_imp_le)
+  next
+    fix x::'a and n
+    assume "x$$i \<le> a - 1 / real (Suc n)"
+    also have "\<dots> < a" by auto
+    finally show "x$$i < a" .
+  qed
+  show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: countable_UN)
+       (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_span_halfspace_ge:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq> 
+   sets (sigma UNIV (range (\<lambda> (a, i). {x. a \<le> x $$ i})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
+  show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: Diff)
+       (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_le_span_halfspace_gt:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq> 
+   sets (sigma UNIV (range (\<lambda> (a, i). {x. a < x $$ i})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+  show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: Diff)
+       (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_le_span_atMost:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
+   sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a i
+  show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+  proof cases
+    assume "i < DIM('a)"
+    then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
+    proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
+      fix x
+      from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
+      then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
+        by (subst (asm) Max_le_iff) auto
+      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
+        by (auto intro!: exI[of _ k])
+    qed
+    show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: countable_UN)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  next
+    assume "\<not> i < DIM('a)"
+    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      using top by auto
+  qed
+qed auto
+
+lemma halfspace_le_span_greaterThan:
+  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
+   sets (sigma UNIV (range (\<lambda>a. {a<..})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a i
+  show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+  proof cases
+    assume "i < DIM('a)"
+    have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+    also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
+    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+      fix x
+      from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
+      guess k::nat .. note k = this
+      { fix i assume "i < DIM('a)"
+        then have "-x$$i < real k"
+          using k by (subst (asm) Max_less_iff) auto
+        then have "- real k < x$$i" by simp }
+      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
+        by (auto intro!: exI[of _ k])
+    qed
+    finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      apply (simp only:)
+      apply (safe intro!: countable_UN Diff)
+      by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  next
+    assume "\<not> i < DIM('a)"
+    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      using top by auto
+  qed
+qed auto
+
+lemma atMost_span_atLeastAtMost:
+  "sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))) \<subseteq>
+   sets (sigma UNIV (range (\<lambda>(a,b). {a..b})))"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  fix a::'a
+  have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+  proof (safe, simp_all add: eucl_le[where 'a='a])
+    fix x
+    from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
+    guess k::nat .. note k = this
+    { fix i assume "i < DIM('a)"
+      with k have "- x$$i \<le> real k"
+        by (subst (asm) Max_le_iff) (auto simp: field_simps)
+      then have "- real k \<le> x$$i" by simp }
+    then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
+      by (auto intro!: exI[of _ k])
+  qed
+  show "{..a} \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: countable_UN)
+       (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma borel_space_eq_greaterThanLessThan:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})))"
+    (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    unfolding borel_space_def
+  proof (rule sigma_algebra.sets_sigma_subset, safe)
+    show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+    then interpret sigma_algebra ?SIGMA .
+    fix M :: "'a set" assume "M \<in> open"
+    then have "open M" by (simp add: mem_def)
+    show "M \<in> sets ?SIGMA"
+      apply (subst open_UNION[OF `open M`])
+      apply (safe intro!: countable_UN)
+      by (auto simp add: sigma_def intro!: sigma_sets.Basic)
+  qed auto
+qed
+
+lemma borel_space_eq_atMost:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})))"
+    (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_atLeastAtMost:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using atMost_span_atLeastAtMost halfspace_le_span_atMost
+      halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_greaterThan:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_greaterThan
+      halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_le:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using open_span_halfspace halfspace_span_halfspace_le by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_less:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using open_span_halfspace .
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_gt:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_ge:
+  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})))"
+   (is "_ = sets ?SIGMA")
+proof (rule antisym)
+  show "sets borel_space \<subseteq> sets ?SIGMA"
+    using halfspace_span_halfspace_ge open_span_halfspace by auto
+  show "sets ?SIGMA \<subseteq> sets borel_space"
+    by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_halfspacesI:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  assumes "sets borel_space = sets (sigma UNIV (range F))"
+  and "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
+  and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
+proof safe
+  fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
+  then show "S a i \<in> sets M" unfolding assms
+    by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def)
+next
+  assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
+  { fix a i have "S a i \<in> sets M"
+    proof cases
+      assume "i < DIM('c)"
+      with a show ?thesis unfolding assms(2) by simp
+    next
+      assume "\<not> i < DIM('c)"
+      from assms(3)[OF this] show ?thesis .
+    qed }
+  then have "f \<in> measurable M (sigma UNIV (range F))"
+    by (auto intro!: measurable_sigma simp: assms(2))
+  then show "f \<in> borel_measurable M" unfolding measurable_def
+    unfolding assms(1) by simp
+qed
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_le]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_less:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_less]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_ge]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_gt]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_le:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
+  using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_less:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
+  using borel_measurable_iff_halfspace_less[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_ge:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
+  using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_greater:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
+  using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
+
+subsection "Borel measurable operators"
+
+lemma (in sigma_algebra) affine_borel_measurable_vector:
+  fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
+proof (rule borel_measurableI)
+  fix S :: "'x set" assume "open S"
+  show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
+  proof cases
+    assume "b \<noteq> 0"
+    with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open")
+      by (auto intro!: open_affinity simp: scaleR.add_right mem_def)
+    hence "?S \<in> sets borel_space"
+      unfolding borel_space_def by (auto simp: sigma_def intro!: sigma_sets.Basic)
+    moreover
+    from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+      apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
+    ultimately show ?thesis using assms unfolding in_borel_measurable_borel_space
+      by auto
+  qed simp
+qed
+
+lemma (in sigma_algebra) affine_borel_measurable:
+  fixes g :: "'a \<Rightarrow> real"
+  assumes g: "g \<in> borel_measurable M"
+  shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
+  using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
+
+lemma (in sigma_algebra) borel_measurable_add[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
 proof -
-  have 1:"!!a. {w \<in> space M. a \<le> f w + g w} = {w \<in> space M. a + (g w) * -1 \<le> f w}"
+  have 1: "\<And>a. {w\<in>space M. a \<le> f w + g w} = {w \<in> space M. a + g w * -1 \<le> f w}"
     by auto
-  have "!!a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"
-    by (rule affine_borel_measurable [OF g]) 
-  hence "!!a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
-    by (rule borel_measurable_leq_borel_measurable) 
-  hence "!!a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
-    by (simp add: 1) 
-  thus ?thesis
-    by (simp add: borel_measurable_ge_iff) 
+  have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"
+    by (rule affine_borel_measurable [OF g])
+  then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
+    by auto
+  then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
+    by (simp add: 1)
+  then show ?thesis
+    by (simp add: borel_measurable_iff_ge)
 qed
 
-
-lemma (in measure_space) borel_measurable_square:
+lemma (in sigma_algebra) borel_measurable_square:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M"
 proof -
@@ -286,21 +715,21 @@
     have "{w \<in> space M. (f w)\<twosuperior> \<le> a} \<in> sets M"
     proof (cases rule: linorder_cases [of a 0])
       case less
-      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}" 
+      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}"
         by auto (metis less order_le_less_trans power2_less_0)
       also have "... \<in> sets M"
-        by (rule empty_sets) 
+        by (rule empty_sets)
       finally show ?thesis .
     next
       case equal
-      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = 
+      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
              {w \<in> space M. f w \<le> 0} \<inter> {w \<in> space M. 0 \<le> f w}"
         by auto
       also have "... \<in> sets M"
-        apply (insert f) 
-        apply (rule Int) 
-        apply (simp add: borel_measurable_le_iff)
-        apply (simp add: borel_measurable_ge_iff)
+        apply (insert f)
+        apply (rule Int)
+        apply (simp add: borel_measurable_iff_le)
+        apply (simp add: borel_measurable_iff_ge)
         done
       finally show ?thesis .
     next
@@ -309,329 +738,536 @@
         by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs
                   real_sqrt_le_iff real_sqrt_power)
       hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
-             {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}" 
+             {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}"
         using greater by auto
       also have "... \<in> sets M"
-        apply (insert f) 
-        apply (rule Int) 
-        apply (simp add: borel_measurable_ge_iff)
-        apply (simp add: borel_measurable_le_iff)
+        apply (insert f)
+        apply (rule Int)
+        apply (simp add: borel_measurable_iff_ge)
+        apply (simp add: borel_measurable_iff_le)
         done
       finally show ?thesis .
     qed
   }
-  thus ?thesis by (auto simp add: borel_measurable_le_iff) 
+  thus ?thesis by (auto simp add: borel_measurable_iff_le)
 qed
 
 lemma times_eq_sum_squares:
    fixes x::real
    shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4"
-by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric]) 
+by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric])
 
-
-lemma (in measure_space) borel_measurable_uminus_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]:
+  fixes g :: "'a \<Rightarrow> real"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. - g x) \<in> borel_measurable M"
 proof -
   have "(\<lambda>x. - g x) = (\<lambda>x. 0 + (g x) * -1)"
     by simp
-  also have "... \<in> borel_measurable M" 
-    by (fast intro: affine_borel_measurable g) 
+  also have "... \<in> borel_measurable M"
+    by (fast intro: affine_borel_measurable g)
   finally show ?thesis .
 qed
 
-lemma (in measure_space) borel_measurable_times_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_times[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
 proof -
   have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M"
-    by (fast intro: affine_borel_measurable borel_measurable_square 
-                    borel_measurable_add_borel_measurable f g) 
-  have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) = 
+    using assms by (fast intro: affine_borel_measurable borel_measurable_square)
+  have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) =
         (\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))"
     by (simp add: minus_divide_right)
-  also have "... \<in> borel_measurable M" 
-    by (fast intro: affine_borel_measurable borel_measurable_square 
-                    borel_measurable_add_borel_measurable 
-                    borel_measurable_uminus_borel_measurable f g)
+  also have "... \<in> borel_measurable M"
+    using f g by (fast intro: affine_borel_measurable borel_measurable_square f g)
   finally have 2: "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) \<in> borel_measurable M" .
   show ?thesis
-    apply (simp add: times_eq_sum_squares diff_minus) 
-    using 1 2 apply (simp add: borel_measurable_add_borel_measurable) 
-    done
+    apply (simp add: times_eq_sum_squares diff_minus)
+    using 1 2 by simp
 qed
 
-lemma (in measure_space) borel_measurable_diff_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_diff[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-unfolding diff_minus
-  by (fast intro: borel_measurable_add_borel_measurable 
-                  borel_measurable_uminus_borel_measurable f g)
+  unfolding diff_minus using assms by fast
 
-lemma (in measure_space) borel_measurable_setsum_borel_measurable:
-  assumes s: "finite s"
-  shows "(!!i. i \<in> s ==> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) s) \<in> borel_measurable M" using s
-proof (induct s)
-  case empty
-  thus ?case
-    by (simp add: borel_measurable_const)
-next
-  case (insert x s)
-  thus ?case
-    by (auto simp add: borel_measurable_add_borel_measurable) 
-qed
+lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
+  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+  assume "finite S"
+  thus ?thesis using assms by induct auto
+qed simp
 
-lemma (in measure_space) borel_measurable_cong:
-  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
-  shows "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M"
-using assms unfolding in_borel_measurable by (simp cong: vimage_inter_cong)
-
-lemma (in measure_space) borel_measurable_inverse:
+lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
-  unfolding borel_measurable_ge_iff
-proof (safe, rule linorder_cases)
-  fix a :: real assume "0 < a"
-  { fix w
-    from `0 < a` have "a \<le> inverse (f w) \<longleftrightarrow> 0 < f w \<and> f w \<le> 1 / a"
-      by (metis inverse_eq_divide inverse_inverse_eq le_imp_inverse_le
-                less_le_trans zero_less_divide_1_iff) }
-  hence "{w \<in> space M. a \<le> inverse (f w)} =
-    {w \<in> space M. 0 < f w} \<inter> {w \<in> space M. f w \<le> 1 / a}" by auto
-  with Int assms[unfolded borel_measurable_gr_iff]
-    assms[unfolded borel_measurable_le_iff]
-  show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
-next
-  fix a :: real assume "0 = a"
-  { fix w have "a \<le> inverse (f w) \<longleftrightarrow> 0 \<le> f w"
-      unfolding `0 = a`[symmetric] by auto }
-  thus "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M"
-    using assms[unfolded borel_measurable_ge_iff] by simp
-next
-  fix a :: real assume "a < 0"
-  { fix w
-    from `a < 0` have "a \<le> inverse (f w) \<longleftrightarrow> f w \<le> 1 / a \<or> 0 \<le> f w"
-      apply (cases "0 \<le> f w")
-      apply (metis inverse_eq_divide linorder_not_le xt1(8) xt1(9)
-                   zero_le_divide_1_iff)
-      apply (metis inverse_eq_divide inverse_inverse_eq inverse_le_imp_le_neg
-                   linorder_not_le order_refl order_trans)
-      done }
-  hence "{w \<in> space M. a \<le> inverse (f w)} =
-    {w \<in> space M. f w \<le> 1 / a} \<union> {w \<in> space M. 0 \<le> f w}" by auto
-  with Un assms[unfolded borel_measurable_ge_iff]
-    assms[unfolded borel_measurable_le_iff]
-  show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
+  unfolding borel_measurable_iff_ge unfolding inverse_eq_divide
+proof safe
+  fix a :: real
+  have *: "{w \<in> space M. a \<le> 1 / f w} =
+      ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
+      ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
+      ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
+  show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
+    by (auto intro!: Int Un)
 qed
 
-lemma (in measure_space) borel_measurable_divide:
+lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   and "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
   unfolding field_divide_inverse
-  by (rule borel_measurable_inverse borel_measurable_times_borel_measurable assms)+
+  by (rule borel_measurable_inverse borel_measurable_times assms)+
+
+lemma (in sigma_algebra) borel_measurable_max[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
+  unfolding borel_measurable_iff_le
+proof safe
+  fix a
+  have "{x \<in> space M. max (g x) (f x) \<le> a} =
+    {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
+  thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
+    using assms unfolding borel_measurable_iff_le
+    by (auto intro!: Int)
+qed
+
+lemma (in sigma_algebra) borel_measurable_min[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
+  unfolding borel_measurable_iff_ge
+proof safe
+  fix a
+  have "{x \<in> space M. a \<le> min (g x) (f x)} =
+    {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
+  thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
+    using assms unfolding borel_measurable_iff_ge
+    by (auto intro!: Int)
+qed
+
+lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
+proof -
+  have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)
+  show ?thesis unfolding * using assms by auto
+qed
+
+section "Borel space over the real line with infinity"
 
-lemma (in measure_space) borel_measurable_vimage:
-  assumes borel: "f \<in> borel_measurable M"
-  shows "f -` {X} \<inter> space M \<in> sets M"
-proof -
-  have "{w \<in> space M. f w = X} = {w. f w = X} \<inter> space M" by auto
-  with borel_measurable_eq_borel_measurable[OF borel borel_measurable_const, of X]
-  show ?thesis unfolding vimage_def by simp
+lemma borel_space_Real_measurable:
+  "A \<in> sets borel_space \<Longrightarrow> Real -` A \<in> sets borel_space"
+proof (rule borel_measurable_translate)
+  fix B :: "pinfreal set" assume "open B"
+  then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
+    x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
+    unfolding open_pinfreal_def by blast
+
+  have "Real -` B = Real -` (B - {\<omega>})" by auto
+  also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
+  also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
+    apply (auto simp add: Real_eq_Real image_iff)
+    apply (rule_tac x="max 0 x" in bexI)
+    by (auto simp: max_def)
+  finally show "Real -` B \<in> sets borel_space"
+    using `open T` by auto
+qed simp
+
+lemma borel_space_real_measurable:
+  "A \<in> sets borel_space \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel_space"
+proof (rule borel_measurable_translate)
+  fix B :: "real set" assume "open B"
+  { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
+  note Ex_less_real = this
+  have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
+    by (force simp: Ex_less_real)
+
+  have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
+    unfolding open_pinfreal_def using `open B`
+    by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
+  then show "(real -` B :: pinfreal set) \<in> sets borel_space" unfolding * by auto
+qed simp
+
+lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+  unfolding in_borel_measurable_borel_space
+proof safe
+  fix S :: "pinfreal set" assume "S \<in> sets borel_space"
+  from borel_space_Real_measurable[OF this]
+  have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
+    using assms
+    unfolding vimage_compose in_borel_measurable_borel_space
+    by auto
+  thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
 qed
 
-section "Monotone convergence"
-
-definition mono_convergent where
-  "mono_convergent u f s \<equiv>
-        (\<forall>x\<in>s. incseq (\<lambda>n. u n x)) \<and>
-        (\<forall>x \<in> s. (\<lambda>i. u i x) ----> f x)"
-
-definition "upclose f g x = max (f x) (g x)"
+lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+  unfolding in_borel_measurable_borel_space
+proof safe
+  fix S :: "real set" assume "S \<in> sets borel_space"
+  from borel_space_real_measurable[OF this]
+  have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
+    using assms
+    unfolding vimage_compose in_borel_measurable_borel_space
+    by auto
+  thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
+qed
 
-primrec mon_upclose where
-"mon_upclose 0 u = u 0" |
-"mon_upclose (Suc n) u = upclose (u (Suc n)) (mon_upclose n u)"
-
-lemma mono_convergentD:
-  assumes "mono_convergent u f s" and "x \<in> s"
-  shows "incseq (\<lambda>n. u n x)" and "(\<lambda>i. u i x) ----> f x"
-  using assms unfolding mono_convergent_def by auto
+lemma (in sigma_algebra) borel_measurable_Real_eq:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+  have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
+    by auto
+  assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+  hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M"
+    by (rule borel_measurable_real)
+  moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x"
+    using assms by auto
+  ultimately show "f \<in> borel_measurable M"
+    by (simp cong: measurable_cong)
+qed auto
 
-lemma mono_convergentI:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> incseq (\<lambda>n. u n x)"
-  assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>i. u i x) ----> f x"
-  shows "mono_convergent u f s"
-  using assms unfolding mono_convergent_def by auto
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real:
+  "f \<in> borel_measurable M \<longleftrightarrow>
+    ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
+proof safe
+  assume "f \<in> borel_measurable M"
+  then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+    by (auto intro: borel_measurable_vimage borel_measurable_real)
+next
+  assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+  have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
+  with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
+  have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))"
+    by (simp add: expand_fun_eq Real_real)
+  show "f \<in> borel_measurable M"
+    apply (subst f)
+    apply (rule measurable_If)
+    using * ** by auto
+qed
+
+lemma (in sigma_algebra) less_eq_ge_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
+proof
+  assume "{x\<in>space M. f x \<le> a} \<in> sets M"
+  moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
+  ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
+next
+  assume "{x\<in>space M. a < f x} \<in> sets M"
+  moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
+  ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
+qed
 
-lemma (in measure_space) mono_convergent_borel_measurable:
-  assumes u: "!!n. u n \<in> borel_measurable M"
-  assumes mc: "mono_convergent u f (space M)"
-  shows "f \<in> borel_measurable M"
-proof -
-  {
-    fix a
-    have 1: "!!w. w \<in> space M & f w <= a \<longleftrightarrow> w \<in> space M & (\<forall>i. u i w <= a)"
+lemma (in sigma_algebra) greater_eq_le_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
+proof
+  assume "{x\<in>space M. a \<le> f x} \<in> sets M"
+  moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
+  ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
+next
+  assume "{x\<in>space M. f x < a} \<in> sets M"
+  moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
+  ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
+qed
+
+lemma (in sigma_algebra) less_eq_le_pinfreal_measurable:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+proof
+  assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
+  show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+  proof
+    fix a show "{x \<in> space M. a < f x} \<in> sets M"
+    proof (cases a)
+      case (preal r)
+      have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
       proof safe
-        fix w i
-        assume w: "w \<in> space M" and f: "f w \<le> a"
-        hence "u i w \<le> f w"
-          by (auto intro: SEQ.incseq_le
-                   simp add: mc [unfolded mono_convergent_def])
-        thus "u i w \<le> a" using f
+        fix x assume "a < f x" and [simp]: "x \<in> space M"
+        with ex_pinfreal_inverse_of_nat_Suc_less[of "f x - a"]
+        obtain n where "a + inverse (of_nat (Suc n)) < f x"
+          by (cases "f x", auto simp: pinfreal_minus_order)
+        then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp
+        then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
           by auto
       next
-        fix w 
-        assume w: "w \<in> space M" and u: "\<forall>i. u i w \<le> a"
-        thus "f w \<le> a"
-          by (metis LIMSEQ_le_const2 mc [unfolded mono_convergent_def])
+        fix i x assume [simp]: "x \<in> space M"
+        have "a < a + inverse (of_nat (Suc i))" using preal by auto
+        also assume "a + inverse (of_nat (Suc i)) \<le> f x"
+        finally show "a < f x" .
       qed
-    have "{w \<in> space M. f w \<le> a} = {w \<in> space M. (\<forall>i. u i w <= a)}"
-      by (simp add: 1)
-    also have "... = (\<Inter>i. {w \<in> space M. u i w \<le> a})" 
-      by auto
-    also have "...  \<in> sets M" using u
-      by (auto simp add: borel_measurable_le_iff intro: countable_INT) 
-    finally have "{w \<in> space M. f w \<le> a} \<in> sets M" .
-  }
-  thus ?thesis 
-    by (auto simp add: borel_measurable_le_iff) 
-qed
-
-lemma mono_convergent_le:
-  assumes "mono_convergent u f s" and "t \<in> s"
-  shows "u n t \<le> f t"
-using mono_convergentD[OF assms] by (auto intro!: incseq_le)
-
-lemma mon_upclose_ex:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ('b\<Colon>linorder)"
-  shows "\<exists>n \<le> m. mon_upclose m u x = u n x"
-proof (induct m)
-  case (Suc m)
-  then obtain n where "n \<le> m" and *: "mon_upclose m u x = u n x" by blast
-  thus ?case
-  proof (cases "u n x \<le> u (Suc m) x")
-    case True with min_max.sup_absorb1 show ?thesis
-      by (auto simp: * upclose_def intro!: exI[of _ "Suc m"])
-  next
-    case False
-     with min_max.sup_absorb2 `n \<le> m` show ?thesis
-       by (auto simp: * upclose_def intro!: exI[of _ n] min_max.sup_absorb2)
+      with a show ?thesis by auto
+    qed simp
   qed
-qed simp
-
-lemma mon_upclose_all:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ('b\<Colon>linorder)"
-  assumes "m \<le> n"
-  shows "u m x \<le> mon_upclose n u x"
-using assms proof (induct n)
-  case 0 thus ?case by auto
 next
-  case (Suc n)
-  show ?case
-  proof (cases "m = Suc n")
-    case True thus ?thesis by (simp add: upclose_def)
-  next
-    case False
-    hence "m \<le> n" using `m \<le> Suc n` by simp
-    from Suc.hyps[OF this]
-    show ?thesis by (auto simp: upclose_def intro!: min_max.le_supI2)
+  assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+  then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
+  show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
+  proof
+    fix a show "{x \<in> space M. f x < a} \<in> sets M"
+    proof (cases a)
+      case (preal r)
+      show ?thesis
+      proof cases
+        assume "a = 0" then show ?thesis by simp
+      next
+        assume "a \<noteq> 0"
+        have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
+        proof safe
+          fix x assume "f x < a" and [simp]: "x \<in> space M"
+          with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"]
+          obtain n where "inverse (of_nat (Suc n)) < a - f x"
+            using preal by (cases "f x") auto
+          then have "f x \<le> a - inverse (of_nat (Suc n)) "
+            using preal by (cases "f x") (auto split: split_if_asm)
+          then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
+            by auto
+        next
+          fix i x assume [simp]: "x \<in> space M"
+          assume "f x \<le> a - inverse (of_nat (Suc i))"
+          also have "\<dots> < a" using `a \<noteq> 0` preal by auto
+          finally show "f x < a" .
+        qed
+        with a show ?thesis by auto
+      qed
+    next
+      case infinite
+      have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
+      proof (safe, simp_all, safe)
+        fix x assume *: "\<forall>n::nat. Real (real n) < f x"
+        show "f x = \<omega>"    proof (rule ccontr)
+          assume "f x \<noteq> \<omega>"
+          with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
+            by (auto simp: pinfreal_noteq_omega_Ex)
+          with *[THEN spec, of n] show False by auto
+        qed
+      qed
+      with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
+      moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
+        using infinite by auto
+      ultimately show ?thesis by auto
+    qed
   qed
 qed
 
-lemma mono_convergent_limit:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes "mono_convergent u f s" and "x \<in> s" and "0 < r"
-  shows "\<exists>N. \<forall>n\<ge>N. f x - u n x < r"
-proof -
-  from LIMSEQ_D[OF mono_convergentD(2)[OF assms(1,2)] `0 < r`]
-  obtain N where "\<And>n. N \<le> n \<Longrightarrow> \<bar> u n x - f x \<bar> < r" by auto
-  with mono_convergent_le[OF assms(1,2)] `0 < r`
-  show ?thesis by (auto intro!: exI[of _ N])
-qed
-
-lemma mon_upclose_le_mono_convergent:
-  assumes mc: "\<And>n. mono_convergent (\<lambda>m. u m n) (f n) s" and "x \<in> s"
-  and "incseq (\<lambda>n. f n x)"
-  shows "mon_upclose n (u n) x <= f n x"
-proof -
-  obtain m where *: "mon_upclose n (u n) x = u n m x" and "m \<le> n"
-    using mon_upclose_ex[of n "u n" x] by auto
-  note this(1)
-  also have "u n m x \<le> f m x" using mono_convergent_le[OF assms(1,2)] .
-  also have "... \<le> f n x" using assms(3) `m \<le> n` unfolding incseq_def by auto
-  finally show ?thesis .
-qed
-
-lemma mon_upclose_mono_convergent:
-  assumes mc_u: "\<And>n. mono_convergent (\<lambda>m. u m n) (f n) s"
-  and mc_f: "mono_convergent f h s"
-  shows "mono_convergent (\<lambda>n. mon_upclose n (u n)) h s"
-proof (rule mono_convergentI)
-  fix x assume "x \<in> s"
-  show "incseq (\<lambda>n. mon_upclose n (u n) x)" unfolding incseq_def
-  proof safe
-    fix m n :: nat assume "m \<le> n"
-    obtain i where mon: "mon_upclose m (u m) x = u m i x" and "i \<le> m"
-      using mon_upclose_ex[of m "u m" x] by auto
-    hence "i \<le> n" using `m \<le> n` by auto
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
+proof safe
+  fix a assume f: "f \<in> borel_measurable M"
+  have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
+  with f show "{x\<in>space M. a < f x} \<in> sets M"
+    by (auto intro!: measurable_sets)
+next
+  assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
+  hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
+    unfolding less_eq_le_pinfreal_measurable
+    unfolding greater_eq_le_measurable .
 
-    note mon
-    also have "u m i x \<le> u n i x"
-      using mono_convergentD(1)[OF mc_u `x \<in> s`] `m \<le> n`
-      unfolding incseq_def by auto
-    also have "u n i x \<le> mon_upclose n (u n) x"
-      using mon_upclose_all[OF `i \<le> n`, of "u n" x] .
-    finally show "mon_upclose m (u m) x \<le> mon_upclose n (u n) x" .
-  qed
-
-  show "(\<lambda>i. mon_upclose i (u i) x) ----> h x"
-  proof (rule LIMSEQ_I)
-    fix r :: real assume "0 < r"
-    hence "0 < r / 2" by auto
-    from mono_convergent_limit[OF mc_f `x \<in> s` this]
-    obtain N where f_h: "\<And>n. N \<le> n \<Longrightarrow> h x - f n x < r / 2" by auto
-
-    from mono_convergent_limit[OF mc_u `x \<in> s` `0 < r / 2`]
-    obtain N' where u_f: "\<And>n. N' \<le> n \<Longrightarrow> f N x - u n N x < r / 2" by auto
+  show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater
+  proof safe
+    have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
+    then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
 
-    show "\<exists>N. \<forall>n\<ge>N. norm (mon_upclose n (u n) x - h x) < r"
-    proof (rule exI[of _ "max N N'"], safe)
-      fix n assume "max N N' \<le> n"
-      hence "N \<le> n" and "N' \<le> n" by auto
-      hence "u n N x \<le> mon_upclose n (u n) x"
-        using mon_upclose_all[of N n "u n" x] by auto
-      moreover
-      from add_strict_mono[OF u_f[OF `N' \<le> n`] f_h[OF order_refl]]
-      have "h x - u n N x < r" by auto
-      ultimately have "h x - mon_upclose n (u n) x < r" by auto
-      moreover
-      obtain i where "mon_upclose n (u n) x = u n i x"
-        using mon_upclose_ex[of n "u n"] by blast
-      with mono_convergent_le[OF mc_u `x \<in> s`, of n i]
-           mono_convergent_le[OF mc_f `x \<in> s`, of i]
-      have "mon_upclose n (u n) x \<le> h x" by auto
-      ultimately
-      show "norm (mon_upclose n (u n) x - h x) < r" by auto
-     qed
+    fix a
+    have "{w \<in> space M. a < real (f w)} =
+      (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
+    proof (split split_if, safe del: notI)
+      fix x assume "0 \<le> a"
+      { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
+          using `0 \<le> a` by (cases "f x", auto) }
+      { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
+          using `0 \<le> a` by (cases "f x", auto) }
+    next
+      fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto
+    qed
+    then show "{w \<in> space M. a < real (f w)} \<in> sets M"
+      using \<omega> * by (auto intro!: Diff)
   qed
 qed
 
-lemma mono_conv_outgrow:
-  assumes "incseq x" "x ----> y" "z < y"
-  shows "\<exists>b. \<forall> a \<ge> b. z < x a"
-using assms
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_le:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_ge:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_const:
+  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M"
+  shows "{x\<in>space M. f x = c} \<in> sets M"
+proof -
+  have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
+  then show ?thesis using assms by (auto intro!: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_neq_const:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
+proof -
+  have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
+  then show ?thesis using assms by (auto intro!: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "{x \<in> space M. f x < g x} \<in> sets M"
+proof -
+  have "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+    "(\<lambda>x. real (g x)) \<in> borel_measurable M"
+    using assms by (auto intro!: borel_measurable_real)
+  from borel_measurable_less[OF this]
+  have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
+  moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
+  moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
+  moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
+  moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
+    ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
+    by (auto simp: real_of_pinfreal_strict_mono_iff)
+  ultimately show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+proof -
+  have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
+  then show ?thesis using g f by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "{w \<in> space M. f w = g w} \<in> sets M"
+proof -
+  have "{x \<in> space M. f x = g x} = {x \<in> space M. g x \<le> f x} \<inter> {x \<in> space M. f x \<le> g x}" by auto
+  then show ?thesis using g f by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
 proof -
-  from assms have "y - z > 0" by simp
-  hence A: "\<exists>n. (\<forall> m \<ge> n. \<bar> x m + - y \<bar> < y - z)" using assms
-    unfolding incseq_def LIMSEQ_def dist_real_def diff_minus
-    by simp
-  have "\<forall>m. x m \<le> y" using incseq_le assms by auto
-  hence B: "\<forall>m. \<bar> x m + - y \<bar> = y - x m"
-    by (metis abs_if abs_minus_add_cancel less_iff_diff_less_0 linorder_not_le diff_minus)
-  from A B show ?thesis by auto
+  have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
+  thus ?thesis using f g by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+proof -
+  have *: "(\<lambda>x. f x + g x) =
+     (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
+     by (auto simp: expand_fun_eq pinfreal_noteq_omega_Ex)
+  show ?thesis using assms unfolding *
+    by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+proof -
+  have *: "(\<lambda>x. f x * g x) =
+     (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
+      Real (real (f x) * real (g x)))"
+     by (auto simp: expand_fun_eq pinfreal_noteq_omega_Ex)
+  show ?thesis using assms unfolding *
+    by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+  assume "finite S"
+  thus ?thesis using assms
+    by induct auto
+qed (simp add: borel_measurable_const)
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_min[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
+  using assms unfolding min_def by (auto intro!: measurable_If)
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_max[intro]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  and "g \<in> borel_measurable M"
+  shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
+  using assms unfolding max_def by (auto intro!: measurable_If)
+
+lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
+  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
+  unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+  fix a
+  have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
+    by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'b=pinfreal])
+  then show "{x\<in>space M. a < ?sup x} \<in> sets M"
+    using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
+  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
+  unfolding borel_measurable_pinfreal_iff_less
+proof safe
+  fix a
+  have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
+    by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand)
+  then show "{x\<in>space M. ?inf x < a} \<in> sets M"
+    using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_diff:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+  unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+  fix a
+  have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
+    by (simp add: pinfreal_less_minus_iff)
+  then show "{x \<in> space M. a < f x - g x} \<in> sets M"
+    using assms by auto
 qed
 
 end