merged
authorhoelzl
Mon, 14 Mar 2011 15:29:10 +0100
changeset 41982 96cbc6379e5a
parent 41968 7f5c9bd991be (current diff)
parent 41981 cdf7693bbe08 (diff)
child 41983 2dc6e382a58b
merged
src/HOL/Probability/Positive_Extended_Real.thy
--- a/src/HOL/Complete_Lattice.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Complete_Lattice.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -89,25 +89,45 @@
   by (auto intro: Sup_least dest: Sup_upper)
 
 lemma Inf_mono:
-  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
-  shows "Inf A \<le> Inf B"
+  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
+  shows "Inf A \<sqsubseteq> Inf B"
 proof (rule Inf_greatest)
   fix b assume "b \<in> B"
-  with assms obtain a where "a \<in> A" and "a \<le> b" by blast
-  from `a \<in> A` have "Inf A \<le> a" by (rule Inf_lower)
-  with `a \<le> b` show "Inf A \<le> b" by auto
+  with assms obtain a where "a \<in> A" and "a \<sqsubseteq> b" by blast
+  from `a \<in> A` have "Inf A \<sqsubseteq> a" by (rule Inf_lower)
+  with `a \<sqsubseteq> b` show "Inf A \<sqsubseteq> b" by auto
 qed
 
 lemma Sup_mono:
-  assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
-  shows "Sup A \<le> Sup B"
+  assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<sqsubseteq> b"
+  shows "Sup A \<sqsubseteq> Sup B"
 proof (rule Sup_least)
   fix a assume "a \<in> A"
-  with assms obtain b where "b \<in> B" and "a \<le> b" by blast
-  from `b \<in> B` have "b \<le> Sup B" by (rule Sup_upper)
-  with `a \<le> b` show "a \<le> Sup B" by auto
+  with assms obtain b where "b \<in> B" and "a \<sqsubseteq> b" by blast
+  from `b \<in> B` have "b \<sqsubseteq> Sup B" by (rule Sup_upper)
+  with `a \<sqsubseteq> b` show "a \<sqsubseteq> Sup B" by auto
 qed
 
+lemma top_le:
+  "top \<sqsubseteq> x \<Longrightarrow> x = top"
+  by (rule antisym) auto
+
+lemma le_bot:
+  "x \<sqsubseteq> bot \<Longrightarrow> x = bot"
+  by (rule antisym) auto
+
+lemma not_less_bot[simp]: "\<not> (x \<sqsubset> bot)"
+  using bot_least[of x] by (auto simp: le_less)
+
+lemma not_top_less[simp]: "\<not> (top \<sqsubset> x)"
+  using top_greatest[of x] by (auto simp: le_less)
+
+lemma Sup_upper2: "u \<in> A \<Longrightarrow> v \<sqsubseteq> u \<Longrightarrow> v \<sqsubseteq> Sup A"
+  using Sup_upper[of u A] by auto
+
+lemma Inf_lower2: "u \<in> A \<Longrightarrow> u \<sqsubseteq> v \<Longrightarrow> Inf A \<sqsubseteq> v"
+  using Inf_lower[of u A] by auto
+
 definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   "INFI A f = \<Sqinter> (f ` A)"
 
@@ -146,15 +166,27 @@
 context complete_lattice
 begin
 
+lemma SUP_cong: "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> SUPR A f = SUPR A g"
+  by (simp add: SUPR_def cong: image_cong)
+
+lemma INF_cong: "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> INFI A f = INFI A g"
+  by (simp add: INFI_def cong: image_cong)
+
 lemma le_SUPI: "i : A \<Longrightarrow> M i \<sqsubseteq> (SUP i:A. M i)"
   by (auto simp add: SUPR_def intro: Sup_upper)
 
+lemma le_SUPI2: "i \<in> A \<Longrightarrow> u \<sqsubseteq> M i \<Longrightarrow> u \<sqsubseteq> (SUP i:A. M i)"
+  using le_SUPI[of i A M] by auto
+
 lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<sqsubseteq> u) \<Longrightarrow> (SUP i:A. M i) \<sqsubseteq> u"
   by (auto simp add: SUPR_def intro: Sup_least)
 
 lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<sqsubseteq> M i"
   by (auto simp add: INFI_def intro: Inf_lower)
 
+lemma INF_leI2: "i \<in> A \<Longrightarrow> M i \<sqsubseteq> u \<Longrightarrow> (INF i:A. M i) \<sqsubseteq> u"
+  using INF_leI[of i A M] by auto
+
 lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<sqsubseteq> M i) \<Longrightarrow> u \<sqsubseteq> (INF i:A. M i)"
   by (auto simp add: INFI_def intro: Inf_greatest)
 
--- a/src/HOL/IsaMakefile	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/IsaMakefile	Mon Mar 14 15:29:10 2011 +0100
@@ -437,10 +437,10 @@
   Library/ContNotDenum.thy Library/Continuity.thy Library/Convex.thy	\
   Library/Countable.thy Library/Diagonalize.thy Library/Dlist.thy	\
   Library/Efficient_Nat.thy Library/Eval_Witness.thy 			\
-  Library/Executable_Set.thy Library/Float.thy				\
-  Library/Formal_Power_Series.thy Library/Fraction_Field.thy		\
-  Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy		\
-  Library/Function_Algebras.thy						\
+  Library/Executable_Set.thy Library/Extended_Reals.thy			\
+  Library/Float.thy Library/Formal_Power_Series.thy			\
+  Library/Fraction_Field.thy Library/FrechetDeriv.thy Library/Cset.thy	\
+  Library/FuncSet.thy Library/Function_Algebras.thy			\
   Library/Fundamental_Theorem_Algebra.thy Library/Glbs.thy		\
   Library/Indicator_Function.thy Library/Infinite_Set.thy		\
   Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
@@ -619,7 +619,7 @@
   Number_Theory/UniqueFactorization.thy  \
   Number_Theory/ROOT.ML
 	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Number_Theory
-                                     
+
 
 ## HOL-Old_Number_Theory
 
@@ -1154,6 +1154,7 @@
   Multivariate_Analysis/Derivative.thy					\
   Multivariate_Analysis/Determinants.thy				\
   Multivariate_Analysis/Euclidean_Space.thy				\
+  Multivariate_Analysis/Extended_Real_Limits.thy			\
   Multivariate_Analysis/Fashoda.thy					\
   Multivariate_Analysis/Finite_Cartesian_Product.thy			\
   Multivariate_Analysis/Integration.certs				\
@@ -1167,9 +1168,10 @@
   Multivariate_Analysis/Topology_Euclidean_Space.thy			\
   Multivariate_Analysis/document/root.tex				\
   Multivariate_Analysis/normarith.ML Library/Glbs.thy			\
-  Library/Indicator_Function.thy Library/Inner_Product.thy		\
-  Library/Numeral_Type.thy Library/Convex.thy Library/FrechetDeriv.thy	\
-  Library/Product_Vector.thy Library/Product_plus.thy
+  Library/Extended_Reals.thy Library/Indicator_Function.thy		\
+  Library/Inner_Product.thy Library/Numeral_Type.thy Library/Convex.thy	\
+  Library/FrechetDeriv.thy Library/Product_Vector.thy			\
+  Library/Product_plus.thy
 	@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
 
 
@@ -1184,7 +1186,6 @@
   Probability/ex/Koepf_Duermuth_Countermeasure.thy			\
   Probability/Information.thy Probability/Lebesgue_Integration.thy	\
   Probability/Lebesgue_Measure.thy Probability/Measure.thy		\
-  Probability/Positive_Extended_Real.thy				\
   Probability/Probability_Space.thy Probability/Probability.thy		\
   Probability/Product_Measure.thy Probability/Radon_Nikodym.thy		\
   Probability/ROOT.ML Probability/Sigma_Algebra.thy			\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Extended_Reals.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -0,0 +1,2438 @@
+(* Title: src/HOL/Library/Extended_Reals.thy
+   Author: Johannes Hölzl; TU München
+   Author: Robert Himmelmann; TU München
+   Author: Armin Heller; TU München
+   Author: Bogdan Grechuk; University of Edinburgh *)
+
+header {* Extended real number line *}
+
+theory Extended_Reals
+  imports Complex_Main
+begin
+
+text {*
+
+For more lemmas about the extended real numbers go to
+  @{text "src/HOL/Multivaraite_Analysis/Extended_Real_Limits.thy"}
+
+*}
+
+lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
+proof
+  assume "{x..} = UNIV"
+  show "x = bot"
+  proof (rule ccontr)
+    assume "x \<noteq> bot" then have "bot \<notin> {x..}" by (simp add: le_less)
+    then show False using `{x..} = UNIV` by simp
+  qed
+qed auto
+
+lemma SUPR_pair:
+  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
+  by (rule antisym) (auto intro!: SUP_leI le_SUPI2)
+
+lemma INFI_pair:
+  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
+  by (rule antisym) (auto intro!: le_INFI INF_leI2)
+
+subsection {* Definition and basic properties *}
+
+datatype extreal = extreal real | PInfty | MInfty
+
+notation (xsymbols)
+  PInfty  ("\<infinity>")
+
+notation (HTML output)
+  PInfty  ("\<infinity>")
+
+declare [[coercion "extreal :: real \<Rightarrow> extreal"]]
+
+instantiation extreal :: uminus
+begin
+  fun uminus_extreal where
+    "- (extreal r) = extreal (- r)"
+  | "- \<infinity> = MInfty"
+  | "- MInfty = \<infinity>"
+  instance ..
+end
+
+lemma inj_extreal[simp]: "inj_on extreal A"
+  unfolding inj_on_def by auto
+
+lemma MInfty_neq_PInfty[simp]:
+  "\<infinity> \<noteq> - \<infinity>" "- \<infinity> \<noteq> \<infinity>" by simp_all
+
+lemma MInfty_neq_extreal[simp]:
+  "extreal r \<noteq> - \<infinity>" "- \<infinity> \<noteq> extreal r" by simp_all
+
+lemma MInfinity_cases[simp]:
+  "(case - \<infinity> of extreal r \<Rightarrow> f r | \<infinity> \<Rightarrow> y | MInfinity \<Rightarrow> z) = z"
+  by simp
+
+lemma extreal_uminus_uminus[simp]:
+  fixes a :: extreal shows "- (- a) = a"
+  by (cases a) simp_all
+
+lemma MInfty_eq[simp]:
+  "MInfty = - \<infinity>" by simp
+
+declare uminus_extreal.simps(2)[simp del]
+
+lemma extreal_cases[case_names real PInf MInf, cases type: extreal]:
+  assumes "\<And>r. x = extreal r \<Longrightarrow> P"
+  assumes "x = \<infinity> \<Longrightarrow> P"
+  assumes "x = -\<infinity> \<Longrightarrow> P"
+  shows P
+  using assms by (cases x) auto
+
+lemmas extreal2_cases = extreal_cases[case_product extreal_cases]
+lemmas extreal3_cases = extreal2_cases[case_product extreal_cases]
+
+lemma extreal_uminus_eq_iff[simp]:
+  fixes a b :: extreal shows "-a = -b \<longleftrightarrow> a = b"
+  by (cases rule: extreal2_cases[of a b]) simp_all
+
+function of_extreal :: "extreal \<Rightarrow> real" where
+"of_extreal (extreal r) = r" |
+"of_extreal \<infinity> = 0" |
+"of_extreal (-\<infinity>) = 0"
+  by (auto intro: extreal_cases)
+termination proof qed (rule wf_empty)
+
+defs (overloaded)
+  real_of_extreal_def [code_unfold]: "real \<equiv> of_extreal"
+
+lemma real_of_extreal[simp]:
+    "real (- x :: extreal) = - (real x)"
+    "real (extreal r) = r"
+    "real \<infinity> = 0"
+  by (cases x) (simp_all add: real_of_extreal_def)
+
+lemma range_extreal[simp]: "range extreal = UNIV - {\<infinity>, -\<infinity>}"
+proof safe
+  fix x assume "x \<notin> range extreal" "x \<noteq> \<infinity>"
+  then show "x = -\<infinity>" by (cases x) auto
+qed auto
+
+lemma extreal_range_uminus[simp]: "range uminus = (UNIV::extreal set)"
+proof safe
+  fix x :: extreal show "x \<in> range uminus" by (intro image_eqI[of _ _ "-x"]) auto
+qed auto
+
+instantiation extreal :: number
+begin
+definition [simp]: "number_of x = extreal (number_of x)"
+instance proof qed
+end
+
+instantiation extreal :: abs
+begin
+  function abs_extreal where
+    "\<bar>extreal r\<bar> = extreal \<bar>r\<bar>"
+  | "\<bar>-\<infinity>\<bar> = \<infinity>"
+  | "\<bar>\<infinity>\<bar> = \<infinity>"
+  by (auto intro: extreal_cases)
+  termination proof qed (rule wf_empty)
+  instance ..
+end
+
+lemma abs_eq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> = \<infinity> ; x = \<infinity> \<Longrightarrow> P ; x = -\<infinity> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+  by (cases x) auto
+
+lemma abs_neq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> \<noteq> \<infinity> ; \<And>r. x = extreal r \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+  by (cases x) auto
+
+lemma abs_extreal_uminus[simp]: "\<bar>- x\<bar> = \<bar>x::extreal\<bar>"
+  by (cases x) auto
+
+subsubsection "Addition"
+
+instantiation extreal :: comm_monoid_add
+begin
+
+definition "0 = extreal 0"
+
+function plus_extreal where
+"extreal r + extreal p = extreal (r + p)" |
+"\<infinity> + a = \<infinity>" |
+"a + \<infinity> = \<infinity>" |
+"extreal r + -\<infinity> = - \<infinity>" |
+"-\<infinity> + extreal p = -\<infinity>" |
+"-\<infinity> + -\<infinity> = -\<infinity>"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a, b)" by (cases x) auto
+  ultimately show P
+   by (cases rule: extreal2_cases[of a b]) auto
+qed auto
+termination proof qed (rule wf_empty)
+
+lemma Infty_neq_0[simp]:
+  "\<infinity> \<noteq> 0" "0 \<noteq> \<infinity>"
+  "-\<infinity> \<noteq> 0" "0 \<noteq> -\<infinity>"
+  by (simp_all add: zero_extreal_def)
+
+lemma extreal_eq_0[simp]:
+  "extreal r = 0 \<longleftrightarrow> r = 0"
+  "0 = extreal r \<longleftrightarrow> r = 0"
+  unfolding zero_extreal_def by simp_all
+
+instance
+proof
+  fix a :: extreal show "0 + a = a"
+    by (cases a) (simp_all add: zero_extreal_def)
+  fix b :: extreal show "a + b = b + a"
+    by (cases rule: extreal2_cases[of a b]) simp_all
+  fix c :: extreal show "a + b + c = a + (b + c)"
+    by (cases rule: extreal3_cases[of a b c]) simp_all
+qed
+end
+
+lemma abs_extreal_zero[simp]: "\<bar>0\<bar> = (0::extreal)"
+  unfolding zero_extreal_def abs_extreal.simps by simp
+
+lemma extreal_uminus_zero[simp]:
+  "- 0 = (0::extreal)"
+  by (simp add: zero_extreal_def)
+
+lemma extreal_uminus_zero_iff[simp]:
+  fixes a :: extreal shows "-a = 0 \<longleftrightarrow> a = 0"
+  by (cases a) simp_all
+
+lemma extreal_plus_eq_PInfty[simp]:
+  shows "a + b = \<infinity> \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_plus_eq_MInfty[simp]:
+  shows "a + b = -\<infinity> \<longleftrightarrow>
+    (a = -\<infinity> \<or> b = -\<infinity>) \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_add_cancel_left:
+  assumes "a \<noteq> -\<infinity>"
+  shows "a + b = a + c \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
+  using assms by (cases rule: extreal3_cases[of a b c]) auto
+
+lemma extreal_add_cancel_right:
+  assumes "a \<noteq> -\<infinity>"
+  shows "b + a = c + a \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
+  using assms by (cases rule: extreal3_cases[of a b c]) auto
+
+lemma extreal_real:
+  "extreal (real x) = (if \<bar>x\<bar> = \<infinity> then 0 else x)"
+  by (cases x) simp_all
+
+lemma real_of_extreal_add:
+  fixes a b :: extreal
+  shows "real (a + b) = (if (\<bar>a\<bar> = \<infinity>) \<and> (\<bar>b\<bar> = \<infinity>) \<or> (\<bar>a\<bar> \<noteq> \<infinity>) \<and> (\<bar>b\<bar> \<noteq> \<infinity>) then real a + real b else 0)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+subsubsection "Linear order on @{typ extreal}"
+
+instantiation extreal :: linorder
+begin
+
+function less_extreal where
+"extreal x < extreal y \<longleftrightarrow> x < y" |
+"        \<infinity> < a         \<longleftrightarrow> False" |
+"        a < -\<infinity>        \<longleftrightarrow> False" |
+"extreal x < \<infinity>         \<longleftrightarrow> True" |
+"       -\<infinity> < extreal r \<longleftrightarrow> True" |
+"       -\<infinity> < \<infinity>         \<longleftrightarrow> True"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a,b)" by (cases x) auto
+  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
+qed simp_all
+termination by (relation "{}") simp
+
+definition "x \<le> (y::extreal) \<longleftrightarrow> x < y \<or> x = y"
+
+lemma extreal_infty_less[simp]:
+  "x < \<infinity> \<longleftrightarrow> (x \<noteq> \<infinity>)"
+  "-\<infinity> < x \<longleftrightarrow> (x \<noteq> -\<infinity>)"
+  by (cases x, simp_all) (cases x, simp_all)
+
+lemma extreal_infty_less_eq[simp]:
+  "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
+  "x \<le> -\<infinity> \<longleftrightarrow> x = -\<infinity>"
+  by (auto simp add: less_eq_extreal_def)
+
+lemma extreal_less[simp]:
+  "extreal r < 0 \<longleftrightarrow> (r < 0)"
+  "0 < extreal r \<longleftrightarrow> (0 < r)"
+  "0 < \<infinity>"
+  "-\<infinity> < 0"
+  by (simp_all add: zero_extreal_def)
+
+lemma extreal_less_eq[simp]:
+  "x \<le> \<infinity>"
+  "-\<infinity> \<le> x"
+  "extreal r \<le> extreal p \<longleftrightarrow> r \<le> p"
+  "extreal r \<le> 0 \<longleftrightarrow> r \<le> 0"
+  "0 \<le> extreal r \<longleftrightarrow> 0 \<le> r"
+  by (auto simp add: less_eq_extreal_def zero_extreal_def)
+
+lemma extreal_infty_less_eq2:
+  "a \<le> b \<Longrightarrow> a = \<infinity> \<Longrightarrow> b = \<infinity>"
+  "a \<le> b \<Longrightarrow> b = -\<infinity> \<Longrightarrow> a = -\<infinity>"
+  by simp_all
+
+instance
+proof
+  fix x :: extreal show "x \<le> x"
+    by (cases x) simp_all
+  fix y :: extreal show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
+    by (cases rule: extreal2_cases[of x y]) auto
+  show "x \<le> y \<or> y \<le> x "
+    by (cases rule: extreal2_cases[of x y]) auto
+  { assume "x \<le> y" "y \<le> x" then show "x = y"
+    by (cases rule: extreal2_cases[of x y]) auto }
+  { fix z assume "x \<le> y" "y \<le> z" then show "x \<le> z"
+    by (cases rule: extreal3_cases[of x y z]) auto }
+qed
+end
+
+instance extreal :: ordered_ab_semigroup_add
+proof
+  fix a b c :: extreal assume "a \<le> b" then show "c + a \<le> c + b"
+    by (cases rule: extreal3_cases[of a b c]) auto
+qed
+
+lemma extreal_MInfty_lessI[intro, simp]:
+  "a \<noteq> -\<infinity> \<Longrightarrow> -\<infinity> < a"
+  by (cases a) auto
+
+lemma extreal_less_PInfty[intro, simp]:
+  "a \<noteq> \<infinity> \<Longrightarrow> a < \<infinity>"
+  by (cases a) auto
+
+lemma extreal_less_extreal_Ex:
+  fixes a b :: extreal
+  shows "x < extreal r \<longleftrightarrow> x = -\<infinity> \<or> (\<exists>p. p < r \<and> x = extreal p)"
+  by (cases x) auto
+
+lemma less_PInf_Ex_of_nat: "x \<noteq> \<infinity> \<longleftrightarrow> (\<exists>n::nat. x < extreal (real n))"
+proof (cases x)
+  case (real r) then show ?thesis
+    using reals_Archimedean2[of r] by simp
+qed simp_all
+
+lemma extreal_add_mono:
+  fixes a b c d :: extreal assumes "a \<le> b" "c \<le> d" shows "a + c \<le> b + d"
+  using assms
+  apply (cases a)
+  apply (cases rule: extreal3_cases[of b c d], auto)
+  apply (cases rule: extreal3_cases[of b c d], auto)
+  done
+
+lemma extreal_minus_le_minus[simp]:
+  fixes a b :: extreal shows "- a \<le> - b \<longleftrightarrow> b \<le> a"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_minus_less_minus[simp]:
+  fixes a b :: extreal shows "- a < - b \<longleftrightarrow> b < a"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_le_real_iff:
+  "x \<le> real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x \<le> y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x \<le> 0))"
+  by (cases y) auto
+
+lemma real_le_extreal_iff:
+  "real y \<le> x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y \<le> extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 \<le> x))"
+  by (cases y) auto
+
+lemma extreal_less_real_iff:
+  "x < real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x < y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x < 0))"
+  by (cases y) auto
+
+lemma real_less_extreal_iff:
+  "real y < x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y < extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 < x))"
+  by (cases y) auto
+
+lemma real_of_extreal_positive_mono:
+  assumes "x \<noteq> \<infinity>" "y \<noteq> \<infinity>" "0 \<le> x" "x \<le> y"
+  shows "real x \<le> real y"
+  using assms by (cases rule: extreal2_cases[of x y]) auto
+
+lemma real_of_extreal_pos:
+  fixes x :: extreal shows "0 \<le> x \<Longrightarrow> 0 \<le> real x" by (cases x) auto
+
+lemmas real_of_extreal_ord_simps =
+  extreal_le_real_iff real_le_extreal_iff extreal_less_real_iff real_less_extreal_iff
+
+lemma extreal_dense:
+  fixes x y :: extreal assumes "x < y"
+  shows "EX z. x < z & z < y"
+proof -
+{ assume a: "x = (-\<infinity>)"
+  { assume "y = \<infinity>" hence ?thesis using a by (auto intro!: exI[of _ "0"]) }
+  moreover
+  { assume "y ~= \<infinity>"
+    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
+    hence ?thesis using `x < y` a by (auto intro!: exI[of _ "extreal (r - 1)"])
+  } ultimately have ?thesis by auto
+}
+moreover
+{ assume "x ~= (-\<infinity>)"
+  with `x < y` obtain p where p: "x = extreal p" by (cases x) auto
+  { assume "y = \<infinity>" hence ?thesis using `x < y` p
+       by (auto intro!: exI[of _ "extreal (p + 1)"]) }
+  moreover
+  { assume "y ~= \<infinity>"
+    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
+    with p `x < y` have "p < r" by auto
+    with dense obtain z where "p < z" "z < r" by auto
+    hence ?thesis using r p by (auto intro!: exI[of _ "extreal z"])
+  } ultimately have ?thesis by auto
+} ultimately show ?thesis by auto
+qed
+
+lemma extreal_dense2:
+  fixes x y :: extreal assumes "x < y"
+  shows "EX z. x < extreal z & extreal z < y"
+  by (metis extreal_dense[OF `x < y`] extreal_cases less_extreal.simps(2,3))
+
+lemma extreal_add_strict_mono:
+  fixes a b c d :: extreal
+  assumes "a = b" "0 \<le> a" "a \<noteq> \<infinity>" "c < d"
+  shows "a + c < b + d"
+  using assms by (cases rule: extreal3_cases[case_product extreal_cases, of a b c d]) auto
+
+lemma extreal_less_add: "\<bar>a\<bar> \<noteq> \<infinity> \<Longrightarrow> c < b \<Longrightarrow> a + c < a + b"
+  by (cases rule: extreal2_cases[of b c]) auto
+
+lemma extreal_uminus_eq_reorder: "- a = b \<longleftrightarrow> a = (-b::extreal)" by auto
+
+lemma extreal_uminus_less_reorder: "- a < b \<longleftrightarrow> -b < (a::extreal)"
+  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_less_minus)
+
+lemma extreal_uminus_le_reorder: "- a \<le> b \<longleftrightarrow> -b \<le> (a::extreal)"
+  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_le_minus)
+
+lemmas extreal_uminus_reorder =
+  extreal_uminus_eq_reorder extreal_uminus_less_reorder extreal_uminus_le_reorder
+
+lemma extreal_bot:
+  fixes x :: extreal assumes "\<And>B. x \<le> extreal B" shows "x = - \<infinity>"
+proof (cases x)
+  case (real r) with assms[of "r - 1"] show ?thesis by auto
+next case PInf with assms[of 0] show ?thesis by auto
+next case MInf then show ?thesis by simp
+qed
+
+lemma extreal_top:
+  fixes x :: extreal assumes "\<And>B. x \<ge> extreal B" shows "x = \<infinity>"
+proof (cases x)
+  case (real r) with assms[of "r + 1"] show ?thesis by auto
+next case MInf with assms[of 0] show ?thesis by auto
+next case PInf then show ?thesis by simp
+qed
+
+lemma
+  shows extreal_max[simp]: "extreal (max x y) = max (extreal x) (extreal y)"
+    and extreal_min[simp]: "extreal (min x y) = min (extreal x) (extreal y)"
+  by (simp_all add: min_def max_def)
+
+lemma extreal_max_0: "max 0 (extreal r) = extreal (max 0 r)"
+  by (auto simp: zero_extreal_def)
+
+lemma
+  fixes f :: "nat \<Rightarrow> extreal"
+  shows incseq_uminus[simp]: "incseq (\<lambda>x. - f x) \<longleftrightarrow> decseq f"
+  and decseq_uminus[simp]: "decseq (\<lambda>x. - f x) \<longleftrightarrow> incseq f"
+  unfolding decseq_def incseq_def by auto
+
+lemma extreal_add_nonneg_nonneg:
+  fixes a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
+  using add_mono[of 0 a 0 b] by simp
+
+lemma image_eqD: "f ` A = B \<Longrightarrow> (\<forall>x\<in>A. f x \<in> B)"
+  by auto
+
+lemma incseq_setsumI:
+  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
+  assumes "\<And>i. 0 \<le> f i"
+  shows "incseq (\<lambda>i. setsum f {..< i})"
+proof (intro incseq_SucI)
+  fix n have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
+    using assms by (rule add_left_mono)
+  then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
+    by auto
+qed
+
+lemma incseq_setsumI2:
+  fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
+  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
+  shows "incseq (\<lambda>i. \<Sum>n\<in>A. f n i)"
+  using assms unfolding incseq_def by (auto intro: setsum_mono)
+
+subsubsection "Multiplication"
+
+instantiation extreal :: "{comm_monoid_mult, sgn}"
+begin
+
+definition "1 = extreal 1"
+
+function sgn_extreal where
+  "sgn (extreal r) = extreal (sgn r)"
+| "sgn \<infinity> = 1"
+| "sgn (-\<infinity>) = -1"
+by (auto intro: extreal_cases)
+termination proof qed (rule wf_empty)
+
+function times_extreal where
+"extreal r * extreal p = extreal (r * p)" |
+"extreal r * \<infinity> = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
+"\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
+"extreal r * -\<infinity> = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
+"-\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
+"\<infinity> * \<infinity> = \<infinity>" |
+"-\<infinity> * \<infinity> = -\<infinity>" |
+"\<infinity> * -\<infinity> = -\<infinity>" |
+"-\<infinity> * -\<infinity> = \<infinity>"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a, b)" by (cases x) auto
+  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
+qed simp_all
+termination by (relation "{}") simp
+
+instance
+proof
+  fix a :: extreal show "1 * a = a"
+    by (cases a) (simp_all add: one_extreal_def)
+  fix b :: extreal show "a * b = b * a"
+    by (cases rule: extreal2_cases[of a b]) simp_all
+  fix c :: extreal show "a * b * c = a * (b * c)"
+    by (cases rule: extreal3_cases[of a b c])
+       (simp_all add: zero_extreal_def zero_less_mult_iff)
+qed
+end
+
+lemma abs_extreal_one[simp]: "\<bar>1\<bar> = (1::extreal)"
+  unfolding one_extreal_def by simp
+
+lemma extreal_mult_zero[simp]:
+  fixes a :: extreal shows "a * 0 = 0"
+  by (cases a) (simp_all add: zero_extreal_def)
+
+lemma extreal_zero_mult[simp]:
+  fixes a :: extreal shows "0 * a = 0"
+  by (cases a) (simp_all add: zero_extreal_def)
+
+lemma extreal_m1_less_0[simp]:
+  "-(1::extreal) < 0"
+  by (simp add: zero_extreal_def one_extreal_def)
+
+lemma extreal_zero_m1[simp]:
+  "1 \<noteq> (0::extreal)"
+  by (simp add: zero_extreal_def one_extreal_def)
+
+lemma extreal_times_0[simp]:
+  fixes x :: extreal shows "0 * x = 0"
+  by (cases x) (auto simp: zero_extreal_def)
+
+lemma extreal_times[simp]:
+  "1 \<noteq> \<infinity>" "\<infinity> \<noteq> 1"
+  "1 \<noteq> -\<infinity>" "-\<infinity> \<noteq> 1"
+  by (auto simp add: times_extreal_def one_extreal_def)
+
+lemma extreal_plus_1[simp]:
+  "1 + extreal r = extreal (r + 1)" "extreal r + 1 = extreal (r + 1)"
+  "1 + -\<infinity> = -\<infinity>" "-\<infinity> + 1 = -\<infinity>"
+  unfolding one_extreal_def by auto
+
+lemma extreal_zero_times[simp]:
+  fixes a b :: extreal shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_mult_eq_PInfty[simp]:
+  shows "a * b = \<infinity> \<longleftrightarrow>
+    (a = \<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = -\<infinity>)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_mult_eq_MInfty[simp]:
+  shows "a * b = -\<infinity> \<longleftrightarrow>
+    (a = \<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = -\<infinity>)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_0_less_1[simp]: "0 < (1::extreal)"
+  by (simp_all add: zero_extreal_def one_extreal_def)
+
+lemma extreal_zero_one[simp]: "0 \<noteq> (1::extreal)"
+  by (simp_all add: zero_extreal_def one_extreal_def)
+
+lemma extreal_mult_minus_left[simp]:
+  fixes a b :: extreal shows "-a * b = - (a * b)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_mult_minus_right[simp]:
+  fixes a b :: extreal shows "a * -b = - (a * b)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_mult_infty[simp]:
+  "a * \<infinity> = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
+  by (cases a) auto
+
+lemma extreal_infty_mult[simp]:
+  "\<infinity> * a = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
+  by (cases a) auto
+
+lemma extreal_mult_strict_right_mono:
+  assumes "a < b" and "0 < c" "c < \<infinity>"
+  shows "a * c < b * c"
+  using assms
+  by (cases rule: extreal3_cases[of a b c])
+     (auto simp: zero_le_mult_iff extreal_less_PInfty)
+
+lemma extreal_mult_strict_left_mono:
+  "\<lbrakk> a < b ; 0 < c ; c < \<infinity>\<rbrakk> \<Longrightarrow> c * a < c * b"
+  using extreal_mult_strict_right_mono by (simp add: mult_commute[of c])
+
+lemma extreal_mult_right_mono:
+  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a*c \<le> b*c"
+  using assms
+  apply (cases "c = 0") apply simp
+  by (cases rule: extreal3_cases[of a b c])
+     (auto simp: zero_le_mult_iff extreal_less_PInfty)
+
+lemma extreal_mult_left_mono:
+  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> c * a \<le> c * b"
+  using extreal_mult_right_mono by (simp add: mult_commute[of c])
+
+lemma zero_less_one_extreal[simp]: "0 \<le> (1::extreal)"
+  by (simp add: one_extreal_def zero_extreal_def)
+
+lemma extreal_0_le_mult[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * (b :: extreal)"
+  by (cases rule: extreal2_cases[of a b]) (auto simp: mult_nonneg_nonneg)
+
+lemma extreal_right_distrib:
+  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> r * (a + b) = r * a + r * b"
+  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
+
+lemma extreal_left_distrib:
+  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> (a + b) * r = a * r + b * r"
+  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
+
+lemma extreal_mult_le_0_iff:
+  fixes a b :: extreal
+  shows "a * b \<le> 0 \<longleftrightarrow> (0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b)"
+  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_le_0_iff)
+
+lemma extreal_zero_le_0_iff:
+  fixes a b :: extreal
+  shows "0 \<le> a * b \<longleftrightarrow> (0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0)"
+  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_le_mult_iff)
+
+lemma extreal_mult_less_0_iff:
+  fixes a b :: extreal
+  shows "a * b < 0 \<longleftrightarrow> (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b)"
+  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_less_0_iff)
+
+lemma extreal_zero_less_0_iff:
+  fixes a b :: extreal
+  shows "0 < a * b \<longleftrightarrow> (0 < a \<and> 0 < b) \<or> (a < 0 \<and> b < 0)"
+  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_less_mult_iff)
+
+lemma extreal_distrib:
+  fixes a b c :: extreal
+  assumes "a \<noteq> \<infinity> \<or> b \<noteq> -\<infinity>" "a \<noteq> -\<infinity> \<or> b \<noteq> \<infinity>" "\<bar>c\<bar> \<noteq> \<infinity>"
+  shows "(a + b) * c = a * c + b * c"
+  using assms
+  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
+
+lemma extreal_le_epsilon:
+  fixes x y :: extreal
+  assumes "ALL e. 0 < e --> x <= y + e"
+  shows "x <= y"
+proof-
+{ assume a: "EX r. y = extreal r"
+  from this obtain r where r_def: "y = extreal r" by auto
+  { assume "x=(-\<infinity>)" hence ?thesis by auto }
+  moreover
+  { assume "~(x=(-\<infinity>))"
+    from this obtain p where p_def: "x = extreal p"
+    using a assms[rule_format, of 1] by (cases x) auto
+    { fix e have "0 < e --> p <= r + e"
+      using assms[rule_format, of "extreal e"] p_def r_def by auto }
+    hence "p <= r" apply (subst field_le_epsilon) by auto
+    hence ?thesis using r_def p_def by auto
+  } ultimately have ?thesis by blast
+}
+moreover
+{ assume "y=(-\<infinity>) | y=\<infinity>" hence ?thesis
+    using assms[rule_format, of 1] by (cases x) auto
+} ultimately show ?thesis by (cases y) auto
+qed
+
+
+lemma extreal_le_epsilon2:
+  fixes x y :: extreal
+  assumes "ALL e. 0 < e --> x <= y + extreal e"
+  shows "x <= y"
+proof-
+{ fix e :: extreal assume "e>0"
+  { assume "e=\<infinity>" hence "x<=y+e" by auto }
+  moreover
+  { assume "e~=\<infinity>"
+    from this obtain r where "e = extreal r" using `e>0` apply (cases e) by auto
+    hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto
+  } ultimately have "x<=y+e" by blast
+} from this show ?thesis using extreal_le_epsilon by auto
+qed
+
+lemma extreal_le_real:
+  fixes x y :: extreal
+  assumes "ALL z. x <= extreal z --> y <= extreal z"
+  shows "y <= x"
+by (metis assms extreal.exhaust extreal_bot extreal_less_eq(1)
+          extreal_less_eq(2) order_refl uminus_extreal.simps(2))
+
+lemma extreal_le_extreal:
+  fixes x y :: extreal
+  assumes "\<And>B. B < x \<Longrightarrow> B <= y"
+  shows "x <= y"
+by (metis assms extreal_dense leD linorder_le_less_linear)
+
+lemma extreal_ge_extreal:
+  fixes x y :: extreal
+  assumes "ALL B. B>x --> B >= y"
+  shows "x >= y"
+by (metis assms extreal_dense leD linorder_le_less_linear)
+
+subsubsection {* Power *}
+
+instantiation extreal :: power
+begin
+primrec power_extreal where
+  "power_extreal x 0 = 1" |
+  "power_extreal x (Suc n) = x * x ^ n"
+instance ..
+end
+
+lemma extreal_power[simp]: "(extreal x) ^ n = extreal (x^n)"
+  by (induct n) (auto simp: one_extreal_def)
+
+lemma extreal_power_PInf[simp]: "\<infinity> ^ n = (if n = 0 then 1 else \<infinity>)"
+  by (induct n) (auto simp: one_extreal_def)
+
+lemma extreal_power_uminus[simp]:
+  fixes x :: extreal
+  shows "(- x) ^ n = (if even n then x ^ n else - (x^n))"
+  by (induct n) (auto simp: one_extreal_def)
+
+lemma extreal_power_number_of[simp]:
+  "(number_of num :: extreal) ^ n = extreal (number_of num ^ n)"
+  by (induct n) (auto simp: one_extreal_def)
+
+lemma zero_le_power_extreal[simp]:
+  fixes a :: extreal assumes "0 \<le> a"
+  shows "0 \<le> a ^ n"
+  using assms by (induct n) (auto simp: extreal_zero_le_0_iff)
+
+subsubsection {* Subtraction *}
+
+lemma extreal_minus_minus_image[simp]:
+  fixes S :: "extreal set"
+  shows "uminus ` uminus ` S = S"
+  by (auto simp: image_iff)
+
+lemma extreal_uminus_lessThan[simp]:
+  fixes a :: extreal shows "uminus ` {..<a} = {-a<..}"
+proof (safe intro!: image_eqI)
+  fix x assume "-a < x"
+  then have "- x < - (- a)" by (simp del: extreal_uminus_uminus)
+  then show "- x < a" by simp
+qed auto
+
+lemma extreal_uminus_greaterThan[simp]:
+  "uminus ` {(a::extreal)<..} = {..<-a}"
+  by (metis extreal_uminus_lessThan extreal_uminus_uminus
+            extreal_minus_minus_image)
+
+instantiation extreal :: minus
+begin
+definition "x - y = x + -(y::extreal)"
+instance ..
+end
+
+lemma extreal_minus[simp]:
+  "extreal r - extreal p = extreal (r - p)"
+  "-\<infinity> - extreal r = -\<infinity>"
+  "extreal r - \<infinity> = -\<infinity>"
+  "\<infinity> - x = \<infinity>"
+  "-\<infinity> - \<infinity> = -\<infinity>"
+  "x - -y = x + y"
+  "x - 0 = x"
+  "0 - x = -x"
+  by (simp_all add: minus_extreal_def)
+
+lemma extreal_x_minus_x[simp]:
+  "x - x = (if \<bar>x\<bar> = \<infinity> then \<infinity> else 0)"
+  by (cases x) simp_all
+
+lemma extreal_eq_minus_iff:
+  fixes x y z :: extreal
+  shows "x = z - y \<longleftrightarrow>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y = z) \<and>
+    (y = -\<infinity> \<longrightarrow> x = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> z = \<infinity> \<longrightarrow> x = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>)"
+  by (cases rule: extreal3_cases[of x y z]) auto
+
+lemma extreal_eq_minus:
+  fixes x y z :: extreal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x = z - y \<longleftrightarrow> x + y = z"
+  by (auto simp: extreal_eq_minus_iff)
+
+lemma extreal_less_minus_iff:
+  fixes x y z :: extreal
+  shows "x < z - y \<longleftrightarrow>
+    (y = \<infinity> \<longrightarrow> z = \<infinity> \<and> x \<noteq> \<infinity>) \<and>
+    (y = -\<infinity> \<longrightarrow> x \<noteq> \<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity>\<longrightarrow> x + y < z)"
+  by (cases rule: extreal3_cases[of x y z]) auto
+
+lemma extreal_less_minus:
+  fixes x y z :: extreal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x < z - y \<longleftrightarrow> x + y < z"
+  by (auto simp: extreal_less_minus_iff)
+
+lemma extreal_le_minus_iff:
+  fixes x y z :: extreal
+  shows "x \<le> z - y \<longleftrightarrow>
+    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y \<le> z)"
+  by (cases rule: extreal3_cases[of x y z]) auto
+
+lemma extreal_le_minus:
+  fixes x y z :: extreal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x \<le> z - y \<longleftrightarrow> x + y \<le> z"
+  by (auto simp: extreal_le_minus_iff)
+
+lemma extreal_minus_less_iff:
+  fixes x y z :: extreal
+  shows "x - y < z \<longleftrightarrow>
+    y \<noteq> -\<infinity> \<and> (y = \<infinity> \<longrightarrow> x \<noteq> \<infinity> \<and> z \<noteq> -\<infinity>) \<and>
+    (y \<noteq> \<infinity> \<longrightarrow> x < z + y)"
+  by (cases rule: extreal3_cases[of x y z]) auto
+
+lemma extreal_minus_less:
+  fixes x y z :: extreal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y < z \<longleftrightarrow> x < z + y"
+  by (auto simp: extreal_minus_less_iff)
+
+lemma extreal_minus_le_iff:
+  fixes x y z :: extreal
+  shows "x - y \<le> z \<longleftrightarrow>
+    (y = -\<infinity> \<longrightarrow> z = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> x = \<infinity> \<longrightarrow> z = \<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x \<le> z + y)"
+  by (cases rule: extreal3_cases[of x y z]) auto
+
+lemma extreal_minus_le:
+  fixes x y z :: extreal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y \<le> z \<longleftrightarrow> x \<le> z + y"
+  by (auto simp: extreal_minus_le_iff)
+
+lemma extreal_minus_eq_minus_iff:
+  fixes a b c :: extreal
+  shows "a - b = a - c \<longleftrightarrow>
+    b = c \<or> a = \<infinity> \<or> (a = -\<infinity> \<and> b \<noteq> -\<infinity> \<and> c \<noteq> -\<infinity>)"
+  by (cases rule: extreal3_cases[of a b c]) auto
+
+lemma extreal_add_le_add_iff:
+  "c + a \<le> c + b \<longleftrightarrow>
+    a \<le> b \<or> c = \<infinity> \<or> (c = -\<infinity> \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>)"
+  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
+
+lemma extreal_mult_le_mult_iff:
+  "\<bar>c\<bar> \<noteq> \<infinity> \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+  by (cases rule: extreal3_cases[of a b c]) (simp_all add: mult_le_cancel_left)
+
+lemma extreal_minus_mono:
+  fixes A B C D :: extreal assumes "A \<le> B" "D \<le> C"
+  shows "A - C \<le> B - D"
+  using assms
+  by (cases rule: extreal3_cases[case_product extreal_cases, of A B C D]) simp_all
+
+lemma real_of_extreal_minus:
+  "real (a - b) = (if \<bar>a\<bar> = \<infinity> \<or> \<bar>b\<bar> = \<infinity> then 0 else real a - real b)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_diff_positive:
+  fixes a b :: extreal shows "a \<le> b \<Longrightarrow> 0 \<le> b - a"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_between:
+  fixes x e :: extreal
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>" "0 < e"
+  shows "x - e < x" "x < x + e"
+using assms apply (cases x, cases e) apply auto
+using assms by (cases x, cases e) auto
+
+subsubsection {* Division *}
+
+instantiation extreal :: inverse
+begin
+
+function inverse_extreal where
+"inverse (extreal r) = (if r = 0 then \<infinity> else extreal (inverse r))" |
+"inverse \<infinity> = 0" |
+"inverse (-\<infinity>) = 0"
+  by (auto intro: extreal_cases)
+termination by (relation "{}") simp
+
+definition "x / y = x * inverse (y :: extreal)"
+
+instance proof qed
+end
+
+lemma extreal_inverse[simp]:
+  "inverse 0 = \<infinity>"
+  "inverse (1::extreal) = 1"
+  by (simp_all add: one_extreal_def zero_extreal_def)
+
+lemma extreal_divide[simp]:
+  "extreal r / extreal p = (if p = 0 then extreal r * \<infinity> else extreal (r / p))"
+  unfolding divide_extreal_def by (auto simp: divide_real_def)
+
+lemma extreal_divide_same[simp]:
+  "x / x = (if \<bar>x\<bar> = \<infinity> \<or> x = 0 then 0 else 1)"
+  by (cases x)
+     (simp_all add: divide_real_def divide_extreal_def one_extreal_def)
+
+lemma extreal_inv_inv[simp]:
+  "inverse (inverse x) = (if x \<noteq> -\<infinity> then x else \<infinity>)"
+  by (cases x) auto
+
+lemma extreal_inverse_minus[simp]:
+  "inverse (- x) = (if x = 0 then \<infinity> else -inverse x)"
+  by (cases x) simp_all
+
+lemma extreal_uminus_divide[simp]:
+  fixes x y :: extreal shows "- x / y = - (x / y)"
+  unfolding divide_extreal_def by simp
+
+lemma extreal_divide_Infty[simp]:
+  "x / \<infinity> = 0" "x / -\<infinity> = 0"
+  unfolding divide_extreal_def by simp_all
+
+lemma extreal_divide_one[simp]:
+  "x / 1 = (x::extreal)"
+  unfolding divide_extreal_def by simp
+
+lemma extreal_divide_extreal[simp]:
+  "\<infinity> / extreal r = (if 0 \<le> r then \<infinity> else -\<infinity>)"
+  unfolding divide_extreal_def by simp
+
+lemma zero_le_divide_extreal[simp]:
+  fixes a :: extreal assumes "0 \<le> a" "0 \<le> b"
+  shows "0 \<le> a / b"
+  using assms by (cases rule: extreal2_cases[of a b]) (auto simp: zero_le_divide_iff)
+
+lemma extreal_le_divide_pos:
+  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> x * y \<le> z"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_divide_le_pos:
+  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> z \<le> x * y"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_le_divide_neg:
+  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> z \<le> x * y"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_divide_le_neg:
+  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> x * y \<le> z"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_inverse_antimono_strict:
+  fixes x y :: extreal
+  shows "0 \<le> x \<Longrightarrow> x < y \<Longrightarrow> inverse y < inverse x"
+  by (cases rule: extreal2_cases[of x y]) auto
+
+lemma extreal_inverse_antimono:
+  fixes x y :: extreal
+  shows "0 \<le> x \<Longrightarrow> x <= y \<Longrightarrow> inverse y <= inverse x"
+  by (cases rule: extreal2_cases[of x y]) auto
+
+lemma inverse_inverse_Pinfty_iff[simp]:
+  "inverse x = \<infinity> \<longleftrightarrow> x = 0"
+  by (cases x) auto
+
+lemma extreal_inverse_eq_0:
+  "inverse x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity>"
+  by (cases x) auto
+
+lemma extreal_0_gt_inverse:
+  fixes x :: extreal shows "0 < inverse x \<longleftrightarrow> x \<noteq> \<infinity> \<and> 0 \<le> x"
+  by (cases x) auto
+
+lemma extreal_mult_less_right:
+  assumes "b * a < c * a" "0 < a" "a < \<infinity>"
+  shows "b < c"
+  using assms
+  by (cases rule: extreal3_cases[of a b c])
+     (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
+
+lemma extreal_power_divide:
+  "y \<noteq> 0 \<Longrightarrow> (x / y :: extreal) ^ n = x^n / y^n"
+  by (cases rule: extreal2_cases[of x y])
+     (auto simp: one_extreal_def zero_extreal_def power_divide not_le
+                 power_less_zero_eq zero_le_power_iff)
+
+lemma extreal_le_mult_one_interval:
+  fixes x y :: extreal
+  assumes y: "y \<noteq> -\<infinity>"
+  assumes z: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
+  shows "x \<le> y"
+proof (cases x)
+  case PInf with z[of "1 / 2"] show "x \<le> y" by (simp add: one_extreal_def)
+next
+  case (real r) note r = this
+  show "x \<le> y"
+  proof (cases y)
+    case (real p) note p = this
+    have "r \<le> p"
+    proof (rule field_le_mult_one_interval)
+      fix z :: real assume "0 < z" and "z < 1"
+      with z[of "extreal z"]
+      show "z * r \<le> p" using p r by (auto simp: zero_le_mult_iff one_extreal_def)
+    qed
+    then show "x \<le> y" using p r by simp
+  qed (insert y, simp_all)
+qed simp
+
+subsection "Complete lattice"
+
+instantiation extreal :: lattice
+begin
+definition [simp]: "sup x y = (max x y :: extreal)"
+definition [simp]: "inf x y = (min x y :: extreal)"
+instance proof qed simp_all
+end
+
+instantiation extreal :: complete_lattice
+begin
+
+definition "bot = -\<infinity>"
+definition "top = \<infinity>"
+
+definition "Sup S = (LEAST z. ALL x:S. x <= z :: extreal)"
+definition "Inf S = (GREATEST z. ALL x:S. z <= x :: extreal)"
+
+lemma extreal_complete_Sup:
+  fixes S :: "extreal set" assumes "S \<noteq> {}"
+  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
+proof cases
+  assume "\<exists>x. \<forall>a\<in>S. a \<le> extreal x"
+  then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> extreal y" by auto
+  then have "\<infinity> \<notin> S" by force
+  show ?thesis
+  proof cases
+    assume "S = {-\<infinity>}"
+    then show ?thesis by (auto intro!: exI[of _ "-\<infinity>"])
+  next
+    assume "S \<noteq> {-\<infinity>}"
+    with `S \<noteq> {}` `\<infinity> \<notin> S` obtain x where "x \<in> S - {-\<infinity>}" "x \<noteq> \<infinity>" by auto
+    with y `\<infinity> \<notin> S` have "\<forall>z\<in>real ` (S - {-\<infinity>}). z \<le> y"
+      by (auto simp: real_of_extreal_ord_simps)
+    with reals_complete2[of "real ` (S - {-\<infinity>})"] `x \<in> S - {-\<infinity>}`
+    obtain s where s:
+       "\<forall>y\<in>S - {-\<infinity>}. real y \<le> s" "\<And>z. (\<forall>y\<in>S - {-\<infinity>}. real y \<le> z) \<Longrightarrow> s \<le> z"
+       by auto
+    show ?thesis
+    proof (safe intro!: exI[of _ "extreal s"])
+      fix z assume "z \<in> S" with `\<infinity> \<notin> S` show "z \<le> extreal s"
+      proof (cases z)
+        case (real r)
+        then show ?thesis
+          using s(1)[rule_format, of z] `z \<in> S` `z = extreal r` by auto
+      qed auto
+    next
+      fix z assume *: "\<forall>y\<in>S. y \<le> z"
+      with `S \<noteq> {-\<infinity>}` `S \<noteq> {}` show "extreal s \<le> z"
+      proof (cases z)
+        case (real u)
+        with * have "s \<le> u"
+          by (intro s(2)[of u]) (auto simp: real_of_extreal_ord_simps)
+        then show ?thesis using real by simp
+      qed auto
+    qed
+  qed
+next
+  assume *: "\<not> (\<exists>x. \<forall>a\<in>S. a \<le> extreal x)"
+  show ?thesis
+  proof (safe intro!: exI[of _ \<infinity>])
+    fix y assume **: "\<forall>z\<in>S. z \<le> y"
+    with * show "\<infinity> \<le> y"
+    proof (cases y)
+      case MInf with * ** show ?thesis by (force simp: not_le)
+    qed auto
+  qed simp
+qed
+
+lemma extreal_complete_Inf:
+  fixes S :: "extreal set" assumes "S ~= {}"
+  shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)"
+proof-
+def S1 == "uminus ` S"
+hence "S1 ~= {}" using assms by auto
+from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)"
+   using extreal_complete_Sup[of S1] by auto
+{ fix z assume "ALL y:S. z <= y"
+  hence "ALL y:S1. y <= -z" unfolding S1_def by auto
+  hence "x <= -z" using x_def by auto
+  hence "z <= -x"
+    apply (subst extreal_uminus_uminus[symmetric])
+    unfolding extreal_minus_le_minus . }
+moreover have "(ALL y:S. -x <= y)"
+   using x_def unfolding S1_def
+   apply simp
+   apply (subst (3) extreal_uminus_uminus[symmetric])
+   unfolding extreal_minus_le_minus by simp
+ultimately show ?thesis by auto
+qed
+
+lemma extreal_complete_uminus_eq:
+  fixes S :: "extreal set"
+  shows "(\<forall>y\<in>uminus`S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>uminus`S. y \<le> z) \<longrightarrow> x \<le> z)
+     \<longleftrightarrow> (\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
+  by simp (metis extreal_minus_le_minus extreal_uminus_uminus)
+
+lemma extreal_Sup_uminus_image_eq:
+  fixes S :: "extreal set"
+  shows "Sup (uminus ` S) = - Inf S"
+proof cases
+  assume "S = {}"
+  moreover have "(THE x. All (op \<le> x)) = (-\<infinity>::extreal)"
+    by (rule the_equality) (auto intro!: extreal_bot)
+  moreover have "(SOME x. \<forall>y. y \<le> x) = (\<infinity>::extreal)"
+    by (rule some_equality) (auto intro!: extreal_top)
+  ultimately show ?thesis unfolding Inf_extreal_def Sup_extreal_def
+    Least_def Greatest_def GreatestM_def by simp
+next
+  assume "S \<noteq> {}"
+  with extreal_complete_Sup[of "uminus`S"]
+  obtain x where x: "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
+    unfolding extreal_complete_uminus_eq by auto
+  show "Sup (uminus ` S) = - Inf S"
+    unfolding Inf_extreal_def Greatest_def GreatestM_def
+  proof (intro someI2[of _ _ "\<lambda>x. Sup (uminus`S) = - x"])
+    show "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> -x)"
+      using x .
+    fix x' assume "(\<forall>y\<in>S. x' \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> x')"
+    then have "(\<forall>y\<in>uminus`S. y \<le> - x') \<and> (\<forall>y. (\<forall>z\<in>uminus`S. z \<le> y) \<longrightarrow> - x' \<le> y)"
+      unfolding extreal_complete_uminus_eq by simp
+    then show "Sup (uminus ` S) = -x'"
+      unfolding Sup_extreal_def extreal_uminus_eq_iff
+      by (intro Least_equality) auto
+  qed
+qed
+
+instance
+proof
+  { fix x :: extreal and A
+    show "bot <= x" by (cases x) (simp_all add: bot_extreal_def)
+    show "x <= top" by (simp add: top_extreal_def) }
+
+  { fix x :: extreal and A assume "x : A"
+    with extreal_complete_Sup[of A]
+    obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
+    hence "x <= s" using `x : A` by auto
+    also have "... = Sup A" using s unfolding Sup_extreal_def
+      by (auto intro!: Least_equality[symmetric])
+    finally show "x <= Sup A" . }
+  note le_Sup = this
+
+  { fix x :: extreal and A assume *: "!!z. (z : A ==> z <= x)"
+    show "Sup A <= x"
+    proof (cases "A = {}")
+      case True
+      hence "Sup A = -\<infinity>" unfolding Sup_extreal_def
+        by (auto intro!: Least_equality)
+      thus "Sup A <= x" by simp
+    next
+      case False
+      with extreal_complete_Sup[of A]
+      obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
+      hence "Sup A = s"
+        unfolding Sup_extreal_def by (auto intro!: Least_equality)
+      also have "s <= x" using * s by auto
+      finally show "Sup A <= x" .
+    qed }
+  note Sup_le = this
+
+  { fix x :: extreal and A assume "x \<in> A"
+    with le_Sup[of "-x" "uminus`A"] show "Inf A \<le> x"
+      unfolding extreal_Sup_uminus_image_eq by simp }
+
+  { fix x :: extreal and A assume *: "!!z. (z : A ==> x <= z)"
+    with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
+      unfolding extreal_Sup_uminus_image_eq by force }
+qed
+end
+
+lemma extreal_SUPR_uminus:
+  fixes f :: "'a => extreal"
+  shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
+  unfolding SUPR_def INFI_def
+  using extreal_Sup_uminus_image_eq[of "f`R"]
+  by (simp add: image_image)
+
+lemma extreal_INFI_uminus:
+  fixes f :: "'a => extreal"
+  shows "(INF i : R. -(f i)) = -(SUP i : R. f i)"
+  using extreal_SUPR_uminus[of _ "\<lambda>x. - f x"] by simp
+
+lemma extreal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::extreal set)"
+  using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image)
+
+lemma extreal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: extreal set)"
+  by (auto intro!: inj_onI)
+
+lemma extreal_image_uminus_shift:
+  fixes X Y :: "extreal set" shows "uminus ` X = Y \<longleftrightarrow> X = uminus ` Y"
+proof
+  assume "uminus ` X = Y"
+  then have "uminus ` uminus ` X = uminus ` Y"
+    by (simp add: inj_image_eq_iff)
+  then show "X = uminus ` Y" by (simp add: image_image)
+qed (simp add: image_image)
+
+lemma Inf_extreal_iff:
+  fixes z :: extreal
+  shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x<y) <-> Inf X < y"
+  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
+            order_less_le_trans)
+
+lemma Sup_eq_MInfty:
+  fixes S :: "extreal set" shows "Sup S = -\<infinity> \<longleftrightarrow> S = {} \<or> S = {-\<infinity>}"
+proof
+  assume a: "Sup S = -\<infinity>"
+  with complete_lattice_class.Sup_upper[of _ S]
+  show "S={} \<or> S={-\<infinity>}" by auto
+next
+  assume "S={} \<or> S={-\<infinity>}" then show "Sup S = -\<infinity>"
+    unfolding Sup_extreal_def by (auto intro!: Least_equality)
+qed
+
+lemma Inf_eq_PInfty:
+  fixes S :: "extreal set" shows "Inf S = \<infinity> \<longleftrightarrow> S = {} \<or> S = {\<infinity>}"
+  using Sup_eq_MInfty[of "uminus`S"]
+  unfolding extreal_Sup_uminus_image_eq extreal_image_uminus_shift by simp
+
+lemma Inf_eq_MInfty: "-\<infinity> : S ==> Inf S = -\<infinity>"
+  unfolding Inf_extreal_def
+  by (auto intro!: Greatest_equality)
+
+lemma Sup_eq_PInfty: "\<infinity> : S ==> Sup S = \<infinity>"
+  unfolding Sup_extreal_def
+  by (auto intro!: Least_equality)
+
+lemma extreal_SUPI:
+  fixes x :: extreal
+  assumes "!!i. i : A ==> f i <= x"
+  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
+  shows "(SUP i:A. f i) = x"
+  unfolding SUPR_def Sup_extreal_def
+  using assms by (auto intro!: Least_equality)
+
+lemma extreal_INFI:
+  fixes x :: extreal
+  assumes "!!i. i : A ==> f i >= x"
+  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
+  shows "(INF i:A. f i) = x"
+  unfolding INFI_def Inf_extreal_def
+  using assms by (auto intro!: Greatest_equality)
+
+lemma Sup_extreal_close:
+  fixes e :: extreal
+  assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
+  shows "\<exists>x\<in>S. Sup S - e < x"
+  using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1])
+
+lemma Inf_extreal_close:
+  fixes e :: extreal assumes "\<bar>Inf X\<bar> \<noteq> \<infinity>" "0 < e"
+  shows "\<exists>x\<in>X. x < Inf X + e"
+proof (rule Inf_less_iff[THEN iffD1])
+  show "Inf X < Inf X + e" using assms
+    by (cases e) auto
+qed
+
+lemma Sup_eq_top_iff:
+  fixes A :: "'a::{complete_lattice, linorder} set"
+  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
+proof
+  assume *: "Sup A = top"
+  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
+  proof (intro allI impI)
+    fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
+      unfolding less_Sup_iff by auto
+  qed
+next
+  assume *: "\<forall>x<top. \<exists>i\<in>A. x < i"
+  show "Sup A = top"
+  proof (rule ccontr)
+    assume "Sup A \<noteq> top"
+    with top_greatest[of "Sup A"]
+    have "Sup A < top" unfolding le_less by auto
+    then have "Sup A < Sup A"
+      using * unfolding less_Sup_iff by auto
+    then show False by auto
+  qed
+qed
+
+lemma SUP_eq_top_iff:
+  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
+  shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
+  unfolding SUPR_def Sup_eq_top_iff by auto
+
+lemma SUP_nat_Infty: "(SUP i::nat. extreal (real i)) = \<infinity>"
+proof -
+  { fix x assume "x \<noteq> \<infinity>"
+    then have "\<exists>k::nat. x < extreal (real k)"
+    proof (cases x)
+      case MInf then show ?thesis by (intro exI[of _ 0]) auto
+    next
+      case (real r)
+      moreover obtain k :: nat where "r < real k"
+        using ex_less_of_nat by (auto simp: real_eq_of_nat)
+      ultimately show ?thesis by auto
+    qed simp }
+  then show ?thesis
+    using SUP_eq_top_iff[of UNIV "\<lambda>n::nat. extreal (real n)"]
+    by (auto simp: top_extreal_def)
+qed
+
+lemma extreal_le_Sup:
+  fixes x :: extreal
+  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
+(is "?lhs <-> ?rhs")
+proof-
+{ assume "?rhs"
+  { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
+    from this obtain y where y_def: "(SUP i:A. f i)<y & y<x" using extreal_dense by auto
+    from this obtain i where "i : A & y <= f i" using `?rhs` by auto
+    hence "y <= (SUP i:A. f i)" using le_SUPI[of i A f] by auto
+    hence False using y_def by auto
+  } hence "?lhs" by auto
+}
+moreover
+{ assume "?lhs" hence "?rhs"
+  by (metis Collect_def Collect_mem_eq SUP_leI assms atLeastatMost_empty atLeastatMost_empty_iff
+      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
+} ultimately show ?thesis by auto
+qed
+
+lemma extreal_Inf_le:
+  fixes x :: extreal
+  shows "((INF i:A. f i) <= x) <-> (ALL y. x < y --> (EX i. i : A & f i <= y))"
+(is "?lhs <-> ?rhs")
+proof-
+{ assume "?rhs"
+  { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le)
+    from this obtain y where y_def: "x<y & y<(INF i:A. f i)" using extreal_dense by auto
+    from this obtain i where "i : A & f i <= y" using `?rhs` by auto
+    hence "(INF i:A. f i) <= y" using INF_leI[of i A f] by auto
+    hence False using y_def by auto
+  } hence "?lhs" by auto
+}
+moreover
+{ assume "?lhs" hence "?rhs"
+  by (metis Collect_def Collect_mem_eq le_INFI assms atLeastatMost_empty atLeastatMost_empty_iff
+      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
+} ultimately show ?thesis by auto
+qed
+
+lemma Inf_less:
+  fixes x :: extreal
+  assumes "(INF i:A. f i) < x"
+  shows "EX i. i : A & f i <= x"
+proof(rule ccontr)
+  assume "~ (EX i. i : A & f i <= x)"
+  hence "ALL i:A. f i > x" by auto
+  hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto
+  thus False using assms by auto
+qed
+
+lemma same_INF:
+  assumes "ALL e:A. f e = g e"
+  shows "(INF e:A. f e) = (INF e:A. g e)"
+proof-
+have "f ` A = g ` A" unfolding image_def using assms by auto
+thus ?thesis unfolding INFI_def by auto
+qed
+
+lemma same_SUP:
+  assumes "ALL e:A. f e = g e"
+  shows "(SUP e:A. f e) = (SUP e:A. g e)"
+proof-
+have "f ` A = g ` A" unfolding image_def using assms by auto
+thus ?thesis unfolding SUPR_def by auto
+qed
+
+lemma SUPR_eq:
+  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<le> g j"
+  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<le> f i"
+  shows "(SUP i:A. f i) = (SUP j:B. g j)"
+proof (intro antisym)
+  show "(SUP i:A. f i) \<le> (SUP j:B. g j)"
+    using assms by (metis SUP_leI le_SUPI2)
+  show "(SUP i:B. g i) \<le> (SUP j:A. f j)"
+    using assms by (metis SUP_leI le_SUPI2)
+qed
+
+lemma SUP_extreal_le_addI:
+  assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
+  shows "SUPR UNIV f + y \<le> z"
+proof (cases y)
+  case (real r)
+  then have "\<And>i. f i \<le> z - y" using assms by (simp add: extreal_le_minus_iff)
+  then have "SUPR UNIV f \<le> z - y" by (rule SUP_leI)
+  then show ?thesis using real by (simp add: extreal_le_minus_iff)
+qed (insert assms, auto)
+
+lemma SUPR_extreal_add:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
+  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
+proof (rule extreal_SUPI)
+  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
+  have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
+    unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD)
+  { fix j
+    { fix i
+      have "f i + g j \<le> f i + g (max i j)"
+        using `incseq g`[THEN incseqD] by (rule add_left_mono) auto
+      also have "\<dots> \<le> f (max i j) + g (max i j)"
+        using `incseq f`[THEN incseqD] by (rule add_right_mono) auto
+      also have "\<dots> \<le> y" using * by auto
+      finally have "f i + g j \<le> y" . }
+    then have "SUPR UNIV f + g j \<le> y"
+      using assms(4)[of j] by (intro SUP_extreal_le_addI) auto
+    then have "g j + SUPR UNIV f \<le> y" by (simp add: ac_simps) }
+  then have "SUPR UNIV g + SUPR UNIV f \<le> y"
+    using f by (rule SUP_extreal_le_addI)
+  then show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
+qed (auto intro!: add_mono le_SUPI)
+
+lemma SUPR_extreal_add_pos:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes inc: "incseq f" "incseq g" and pos: "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
+  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
+proof (intro SUPR_extreal_add inc)
+  fix i show "f i \<noteq> -\<infinity>" "g i \<noteq> -\<infinity>" using pos[of i] by auto
+qed
+
+lemma SUPR_extreal_setsum:
+  fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> extreal"
+  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)" and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
+  shows "(SUP i. \<Sum>n\<in>A. f n i) = (\<Sum>n\<in>A. SUPR UNIV (f n))"
+proof cases
+  assume "finite A" then show ?thesis using assms
+    by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_extreal_add_pos)
+qed simp
+
+lemma SUPR_extreal_cmult:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
+  shows "(SUP i. c * f i) = c * SUPR UNIV f"
+proof (rule extreal_SUPI)
+  fix i have "f i \<le> SUPR UNIV f" by (rule le_SUPI) auto
+  then show "c * f i \<le> c * SUPR UNIV f"
+    using `0 \<le> c` by (rule extreal_mult_left_mono)
+next
+  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c * f i \<le> y"
+  show "c * SUPR UNIV f \<le> y"
+  proof cases
+    assume c: "0 < c \<and> c \<noteq> \<infinity>"
+    with * have "SUPR UNIV f \<le> y / c"
+      by (intro SUP_leI) (auto simp: extreal_le_divide_pos)
+    with c show ?thesis
+      by (auto simp: extreal_le_divide_pos)
+  next
+    { assume "c = \<infinity>" have ?thesis
+      proof cases
+        assume "\<forall>i. f i = 0"
+        moreover then have "range f = {0}" by auto
+        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
+      next
+        assume "\<not> (\<forall>i. f i = 0)"
+        then obtain i where "f i \<noteq> 0" by auto
+        with *[of i] `c = \<infinity>` `0 \<le> f i` show ?thesis by (auto split: split_if_asm)
+      qed }
+    moreover assume "\<not> (0 < c \<and> c \<noteq> \<infinity>)"
+    ultimately show ?thesis using * `0 \<le> c` by auto
+  qed
+qed
+
+lemma SUP_PInfty:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes "\<And>n::nat. \<exists>i\<in>A. extreal (real n) \<le> f i"
+  shows "(SUP i:A. f i) = \<infinity>"
+  unfolding SUPR_def Sup_eq_top_iff[where 'a=extreal, unfolded top_extreal_def]
+  apply simp
+proof safe
+  fix x assume "x \<noteq> \<infinity>"
+  show "\<exists>i\<in>A. x < f i"
+  proof (cases x)
+    case PInf with `x \<noteq> \<infinity>` show ?thesis by simp
+  next
+    case MInf with assms[of "0"] show ?thesis by force
+  next
+    case (real r)
+    with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < extreal (real n)" by auto
+    moreover from assms[of n] guess i ..
+    ultimately show ?thesis
+      by (auto intro!: bexI[of _ i])
+  qed
+qed
+
+lemma Sup_countable_SUPR:
+  assumes "A \<noteq> {}"
+  shows "\<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
+proof (cases "Sup A")
+  case (real r)
+  have "\<forall>n::nat. \<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
+  proof
+    fix n ::nat have "\<exists>x\<in>A. Sup A - 1 / extreal (real n) < x"
+      using assms real by (intro Sup_extreal_close) (auto simp: one_extreal_def)
+    then guess x ..
+    then show "\<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
+      by (auto intro!: exI[of _ x] simp: extreal_minus_less_iff)
+  qed
+  from choice[OF this] guess f .. note f = this
+  have "SUPR UNIV f = Sup A"
+  proof (rule extreal_SUPI)
+    fix i show "f i \<le> Sup A" using f
+      by (auto intro!: complete_lattice_class.Sup_upper)
+  next
+    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
+    show "Sup A \<le> y"
+    proof (rule extreal_le_epsilon, intro allI impI)
+      fix e :: extreal assume "0 < e"
+      show "Sup A \<le> y + e"
+      proof (cases e)
+        case (real r)
+        hence "0 < r" using `0 < e` by auto
+        then obtain n ::nat where *: "1 / real n < r" "0 < n"
+          using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
+        have "Sup A \<le> f n + 1 / extreal (real n)" using f[THEN spec, of n] by auto
+        also have "1 / extreal (real n) \<le> e" using real * by (auto simp: one_extreal_def )
+        with bound have "f n + 1 / extreal (real n) \<le> y + e" by (rule add_mono) simp
+        finally show "Sup A \<le> y + e" .
+      qed (insert `0 < e`, auto)
+    qed
+  qed
+  with f show ?thesis by (auto intro!: exI[of _ f])
+next
+  case PInf
+  from `A \<noteq> {}` obtain x where "x \<in> A" by auto
+  show ?thesis
+  proof cases
+    assume "\<infinity> \<in> A"
+    moreover then have "\<infinity> \<le> Sup A" by (intro complete_lattice_class.Sup_upper)
+    ultimately show ?thesis by (auto intro!: exI[of _ "\<lambda>x. \<infinity>"])
+  next
+    assume "\<infinity> \<notin> A"
+    have "\<exists>x\<in>A. 0 \<le> x"
+      by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least extreal_infty_less_eq2 linorder_linear)
+    then obtain x where "x \<in> A" "0 \<le> x" by auto
+    have "\<forall>n::nat. \<exists>f. f \<in> A \<and> x + extreal (real n) \<le> f"
+    proof (rule ccontr)
+      assume "\<not> ?thesis"
+      then have "\<exists>n::nat. Sup A \<le> x + extreal (real n)"
+        by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le)
+      then show False using `x \<in> A` `\<infinity> \<notin> A` PInf
+        by(cases x) auto
+    qed
+    from choice[OF this] guess f .. note f = this
+    have "SUPR UNIV f = \<infinity>"
+    proof (rule SUP_PInfty)
+      fix n :: nat show "\<exists>i\<in>UNIV. extreal (real n) \<le> f i"
+        using f[THEN spec, of n] `0 \<le> x`
+        by (cases rule: extreal2_cases[of "f n" x]) (auto intro!: exI[of _ n])
+    qed
+    then show ?thesis using f PInf by (auto intro!: exI[of _ f])
+  qed
+next
+  case MInf
+  with `A \<noteq> {}` have "A = {-\<infinity>}" by (auto simp: Sup_eq_MInfty)
+  then show ?thesis using MInf by (auto intro!: exI[of _ "\<lambda>x. -\<infinity>"])
+qed
+
+lemma SUPR_countable_SUPR:
+  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
+  using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def)
+
+
+lemma Sup_extreal_cadd:
+  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "Sup ((\<lambda>x. a + x) ` A) = a + Sup A"
+proof (rule antisym)
+  have *: "\<And>a::extreal. \<And>A. Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A"
+    by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
+  then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
+  show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
+  proof (cases a)
+    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
+  next
+    case (real r)
+    then have **: "op + (- a) ` op + a ` A = A"
+      by (auto simp: image_iff ac_simps zero_extreal_def[symmetric])
+    from *[of "-a" "(\<lambda>x. a + x) ` A"] real show ?thesis unfolding **
+      by (cases rule: extreal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto
+  qed (insert `a \<noteq> -\<infinity>`, auto)
+qed
+
+lemma Sup_extreal_cminus:
+  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "Sup ((\<lambda>x. a - x) ` A) = a - Inf A"
+  using Sup_extreal_cadd[of "uminus ` A" a] assms
+  by (simp add: comp_def image_image minus_extreal_def
+                 extreal_Sup_uminus_image_eq)
+
+lemma SUPR_extreal_cminus:
+  fixes A assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "(SUP x:A. a - f x) = a - (INF x:A. f x)"
+  using Sup_extreal_cminus[of "f`A" a] assms
+  unfolding SUPR_def INFI_def image_image by auto
+
+lemma Inf_extreal_cminus:
+  fixes A :: "extreal set" assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
+  shows "Inf ((\<lambda>x. a - x) ` A) = a - Sup A"
+proof -
+  { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto }
+  moreover then have "(\<lambda>x. -a - x)`uminus`A = uminus ` (\<lambda>x. a - x) ` A"
+    by (auto simp: image_image)
+  ultimately show ?thesis
+    using Sup_extreal_cminus[of "uminus ` A" "-a"] assms
+    by (auto simp add: extreal_Sup_uminus_image_eq extreal_Inf_uminus_image_eq)
+qed
+
+lemma INFI_extreal_cminus:
+  fixes A assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
+  shows "(INF x:A. a - f x) = a - (SUP x:A. f x)"
+  using Inf_extreal_cminus[of "f`A" a] assms
+  unfolding SUPR_def INFI_def image_image
+  by auto
+
+subsection "Limits on @{typ extreal}"
+
+subsubsection "Topological space"
+
+instantiation extreal :: topological_space
+begin
+
+definition "open A \<longleftrightarrow> open (extreal -` A)
+       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {extreal x <..} \<subseteq> A))
+       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A))"
+
+lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {extreal x<..} \<subseteq> A)"
+  unfolding open_extreal_def by auto
+
+lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A)"
+  unfolding open_extreal_def by auto
+
+lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{extreal x<..} \<subseteq> A"
+  using open_PInfty[OF assms] by auto
+
+lemma open_MInfty2: assumes "open A" "-\<infinity> \<in> A" obtains x where "{..<extreal x} \<subseteq> A"
+  using open_MInfty[OF assms] by auto
+
+lemma extreal_openE: assumes "open A" obtains x y where
+  "open (extreal -` A)"
+  "\<infinity> \<in> A \<Longrightarrow> {extreal x<..} \<subseteq> A"
+  "-\<infinity> \<in> A \<Longrightarrow> {..<extreal y} \<subseteq> A"
+  using assms open_extreal_def by auto
+
+instance
+proof
+  let ?U = "UNIV::extreal set"
+  show "open ?U" unfolding open_extreal_def
+    by (auto intro!: exI[of _ 0])
+next
+  fix S T::"extreal set" assume "open S" and "open T"
+  from `open S`[THEN extreal_openE] guess xS yS .
+  moreover from `open T`[THEN extreal_openE] guess xT yT .
+  ultimately have
+    "open (extreal -` (S \<inter> T))"
+    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {extreal (max xS xT) <..} \<subseteq> S \<inter> T"
+    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< extreal (min yS yT)} \<subseteq> S \<inter> T"
+    by auto
+  then show "open (S Int T)" unfolding open_extreal_def by blast
+next
+  fix K :: "extreal set set" assume "\<forall>S\<in>K. open S"
+  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (extreal -` S) \<and>
+    (\<infinity> \<in> S \<longrightarrow> {extreal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< extreal y} \<subseteq> S)"
+    by (auto simp: open_extreal_def)
+  then show "open (Union K)" unfolding open_extreal_def
+  proof (intro conjI impI)
+    show "open (extreal -` \<Union>K)"
+      using *[THEN choice] by (auto simp: vimage_Union)
+  qed ((metis UnionE Union_upper subset_trans *)+)
+qed
+end
+
+lemma open_extreal: "open S \<Longrightarrow> open (extreal ` S)"
+  by (auto simp: inj_vimage_image_eq open_extreal_def)
+
+lemma open_extreal_vimage: "open S \<Longrightarrow> open (extreal -` S)"
+  unfolding open_extreal_def by auto
+
+lemma open_extreal_lessThan[intro, simp]: "open {..< a :: extreal}"
+proof -
+  have "\<And>x. extreal -` {..<extreal x} = {..< x}"
+    "extreal -` {..< \<infinity>} = UNIV" "extreal -` {..< -\<infinity>} = {}" by auto
+  then show ?thesis by (cases a) (auto simp: open_extreal_def)
+qed
+
+lemma open_extreal_greaterThan[intro, simp]:
+  "open {a :: extreal <..}"
+proof -
+  have "\<And>x. extreal -` {extreal x<..} = {x<..}"
+    "extreal -` {\<infinity><..} = {}" "extreal -` {-\<infinity><..} = UNIV" by auto
+  then show ?thesis by (cases a) (auto simp: open_extreal_def)
+qed
+
+lemma extreal_open_greaterThanLessThan[intro, simp]: "open {a::extreal <..< b}"
+  unfolding greaterThanLessThan_def by auto
+
+lemma closed_extreal_atLeast[simp, intro]: "closed {a :: extreal ..}"
+proof -
+  have "- {a ..} = {..< a}" by auto
+  then show "closed {a ..}"
+    unfolding closed_def using open_extreal_lessThan by auto
+qed
+
+lemma closed_extreal_atMost[simp, intro]: "closed {.. b :: extreal}"
+proof -
+  have "- {.. b} = {b <..}" by auto
+  then show "closed {.. b}"
+    unfolding closed_def using open_extreal_greaterThan by auto
+qed
+
+lemma closed_extreal_atLeastAtMost[simp, intro]:
+  shows "closed {a :: extreal .. b}"
+  unfolding atLeastAtMost_def by auto
+
+lemma closed_extreal_singleton:
+  "closed {a :: extreal}"
+by (metis atLeastAtMost_singleton closed_extreal_atLeastAtMost)
+
+lemma extreal_open_cont_interval:
+  assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
+  obtains e where "e>0" "{x-e <..< x+e} \<subseteq> S"
+proof-
+  from `open S` have "open (extreal -` S)" by (rule extreal_openE)
+  then obtain e where "0 < e" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> extreal y \<in> S"
+    using assms unfolding open_dist by force
+  show thesis
+  proof (intro that subsetI)
+    show "0 < extreal e" using `0 < e` by auto
+    fix y assume "y \<in> {x - extreal e<..<x + extreal e}"
+    with assms obtain t where "y = extreal t" "dist t (real x) < e"
+      apply (cases y) by (auto simp: dist_real_def)
+    then show "y \<in> S" using e[of t] by auto
+  qed
+qed
+
+lemma extreal_open_cont_interval2:
+  assumes "open S" "x \<in> S" and x: "\<bar>x\<bar> \<noteq> \<infinity>"
+  obtains a b where "a < x" "x < b" "{a <..< b} \<subseteq> S"
+proof-
+  guess e using extreal_open_cont_interval[OF assms] .
+  with that[of "x-e" "x+e"] extreal_between[OF x, of e]
+  show thesis by auto
+qed
+
+instance extreal :: t2_space
+proof
+  fix x y :: extreal assume "x ~= y"
+  let "?P x (y::extreal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
+
+  { fix x y :: extreal assume "x < y"
+    from extreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
+    have "?P x y"
+      apply (rule exI[of _ "{..<z}"])
+      apply (rule exI[of _ "{z<..}"])
+      using z by auto }
+  note * = this
+
+  from `x ~= y`
+  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
+  proof (cases rule: linorder_cases)
+    assume "x = y" with `x ~= y` show ?thesis by simp
+  next assume "x < y" from *[OF this] show ?thesis by auto
+  next assume "y < x" from *[OF this] show ?thesis by auto
+  qed
+qed
+
+subsubsection {* Convergent sequences *}
+
+lemma lim_extreal[simp]:
+  "((\<lambda>n. extreal (f n)) ---> extreal x) net \<longleftrightarrow> (f ---> x) net" (is "?l = ?r")
+proof (intro iffI topological_tendstoI)
+  fix S assume "?l" "open S" "x \<in> S"
+  then show "eventually (\<lambda>x. f x \<in> S) net"
+    using `?l`[THEN topological_tendstoD, OF open_extreal, OF `open S`]
+    by (simp add: inj_image_mem_iff)
+next
+  fix S assume "?r" "open S" "extreal x \<in> S"
+  show "eventually (\<lambda>x. extreal (f x) \<in> S) net"
+    using `?r`[THEN topological_tendstoD, OF open_extreal_vimage, OF `open S`]
+    using `extreal x \<in> S` by auto
+qed
+
+lemma lim_real_of_extreal[simp]:
+  assumes lim: "(f ---> extreal x) net"
+  shows "((\<lambda>x. real (f x)) ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume "open S" "x \<in> S"
+  then have S: "open S" "extreal x \<in> extreal ` S"
+    by (simp_all add: inj_image_mem_iff)
+  have "\<forall>x. f x \<in> extreal ` S \<longrightarrow> real (f x) \<in> S" by auto
+  from this lim[THEN topological_tendstoD, OF open_extreal, OF S]
+  show "eventually (\<lambda>x. real (f x) \<in> S) net"
+    by (rule eventually_mono)
+qed
+
+lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= extreal B)" (is "?l = ?r")
+proof assume ?r show ?l apply(rule topological_tendstoI)
+    unfolding eventually_sequentially
+  proof- fix S assume "open S" "\<infinity> : S"
+    from open_PInfty[OF this] guess B .. note B=this
+    from `?r`[rule_format,of "B+1"] guess N .. note N=this
+    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
+    proof safe case goal1
+      have "extreal B < extreal (B + 1)" by auto
+      also have "... <= f n" using goal1 N by auto
+      finally show ?case using B by fastsimp
+    qed
+  qed
+next assume ?l show ?r
+  proof fix B::real have "open {extreal B<..}" "\<infinity> : {extreal B<..}" by auto
+    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
+    guess N .. note N=this
+    show "EX N. ALL n>=N. extreal B <= f n" apply(rule_tac x=N in exI) using N by auto
+  qed
+qed
+
+
+lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= extreal B)" (is "?l = ?r")
+proof assume ?r show ?l apply(rule topological_tendstoI)
+    unfolding eventually_sequentially
+  proof- fix S assume "open S" "(-\<infinity>) : S"
+    from open_MInfty[OF this] guess B .. note B=this
+    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
+    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
+    proof safe case goal1
+      have "extreal (B - 1) >= f n" using goal1 N by auto
+      also have "... < extreal B" by auto
+      finally show ?case using B by fastsimp
+    qed
+  qed
+next assume ?l show ?r
+  proof fix B::real have "open {..<extreal B}" "(-\<infinity>) : {..<extreal B}" by auto
+    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
+    guess N .. note N=this
+    show "EX N. ALL n>=N. extreal B >= f n" apply(rule_tac x=N in exI) using N by auto
+  qed
+qed
+
+
+lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= extreal B" shows "l ~= \<infinity>"
+proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
+  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
+  guess N .. note N=this[rule_format,OF le_refl]
+  hence "extreal ?B <= extreal B" using assms(2)[of N] by(rule order_trans)
+  hence "extreal ?B < extreal ?B" apply (rule le_less_trans) by auto
+  thus False by auto
+qed
+
+
+lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= extreal B" shows "l ~= (-\<infinity>)"
+proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
+  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
+  guess N .. note N=this[rule_format,OF le_refl]
+  hence "extreal B <= extreal ?B" using assms(2)[of N] order_trans[of "extreal B" "f N" "extreal(B - 1)"] by blast
+  thus False by auto
+qed
+
+
+lemma tendsto_explicit:
+  "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
+  unfolding tendsto_def eventually_sequentially by auto
+
+
+lemma tendsto_obtains_N:
+  assumes "f ----> f0"
+  assumes "open S" "f0 : S"
+  obtains N where "ALL n>=N. f n : S"
+  using tendsto_explicit[of f f0] assms by auto
+
+
+lemma tail_same_limit:
+  fixes X Y N
+  assumes "X ----> L" "ALL n>=N. X n = Y n"
+  shows "Y ----> L"
+proof-
+{ fix S assume "open S" and "L:S"
+  from this obtain N1 where "ALL n>=N1. X n : S"
+     using assms unfolding tendsto_def eventually_sequentially by auto
+  hence "ALL n>=max N N1. Y n : S" using assms by auto
+  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
+}
+thus ?thesis using tendsto_explicit by auto
+qed
+
+
+lemma Lim_bounded_PInfty2:
+assumes lim:"f ----> l" and "ALL n>=N. f n <= extreal B"
+shows "l ~= \<infinity>"
+proof-
+  def g == "(%n. if n>=N then f n else extreal B)"
+  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
+  moreover have "!!n. g n <= extreal B" using g_def assms by auto
+  ultimately show ?thesis using  Lim_bounded_PInfty by auto
+qed
+
+lemma Lim_bounded_extreal:
+  assumes lim:"f ----> (l :: extreal)"
+  and "ALL n>=M. f n <= C"
+  shows "l<=C"
+proof-
+{ assume "l=(-\<infinity>)" hence ?thesis by auto }
+moreover
+{ assume "~(l=(-\<infinity>))"
+  { assume "C=\<infinity>" hence ?thesis by auto }
+  moreover
+  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
+    hence "l=(-\<infinity>)" using assms
+       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
+    hence ?thesis by auto }
+  moreover
+  { assume "EX B. C = extreal B"
+    from this obtain B where B_def: "C=extreal B" by auto
+    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
+    from this obtain m where m_def: "extreal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
+    from this obtain N where N_def: "ALL n>=N. f n : {extreal(m - 1) <..< extreal(m+1)}"
+       apply (subst tendsto_obtains_N[of f l "{extreal(m - 1) <..< extreal(m+1)}"]) using assms by auto
+    { fix n assume "n>=N"
+      hence "EX r. extreal r = f n" using N_def by (cases "f n") auto
+    } from this obtain g where g_def: "ALL n>=N. extreal (g n) = f n" by metis
+    hence "(%n. extreal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
+    hence *: "(%n. g n) ----> m" using m_def by auto
+    { fix n assume "n>=max N M"
+      hence "extreal (g n) <= extreal B" using assms g_def B_def by auto
+      hence "g n <= B" by auto
+    } hence "EX N. ALL n>=N. g n <= B" by blast
+    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
+    hence ?thesis using m_def B_def by auto
+  } ultimately have ?thesis by (cases C) auto
+} ultimately show ?thesis by blast
+qed
+
+lemma real_of_extreal_0[simp]: "real (0::extreal) = 0"
+  unfolding real_of_extreal_def zero_extreal_def by simp
+
+lemma real_of_extreal_mult[simp]:
+  fixes a b :: extreal shows "real (a * b) = real a * real b"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma real_of_extreal_eq_0:
+  "real x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity> \<or> x = 0"
+  by (cases x) auto
+
+lemma tendsto_extreal_realD:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes "x \<noteq> 0" and tendsto: "((\<lambda>x. extreal (real (f x))) ---> x) net"
+  shows "(f ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume S: "open S" "x \<in> S"
+  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}" by auto
+  from tendsto[THEN topological_tendstoD, OF this]
+  show "eventually (\<lambda>x. f x \<in> S) net"
+    by (rule eventually_rev_mp) (auto simp: extreal_real real_of_extreal_0)
+qed
+
+lemma tendsto_extreal_realI:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>" and tendsto: "(f ---> x) net"
+  shows "((\<lambda>x. extreal (real (f x))) ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume "open S" "x \<in> S"
+  with x have "open (S - {\<infinity>, -\<infinity>})" "x \<in> S - {\<infinity>, -\<infinity>}" by auto
+  from tendsto[THEN topological_tendstoD, OF this]
+  show "eventually (\<lambda>x. extreal (real (f x)) \<in> S) net"
+    by (elim eventually_elim1) (auto simp: extreal_real)
+qed
+
+lemma extreal_mult_cancel_left:
+  fixes a b c :: extreal shows "a * b = a * c \<longleftrightarrow>
+    ((\<bar>a\<bar> = \<infinity> \<and> 0 < b * c) \<or> a = 0 \<or> b = c)"
+  by (cases rule: extreal3_cases[of a b c])
+     (simp_all add: zero_less_mult_iff)
+
+lemma extreal_inj_affinity:
+  assumes "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" "\<bar>t\<bar> \<noteq> \<infinity>"
+  shows "inj_on (\<lambda>x. m * x + t) A"
+  using assms
+  by (cases rule: extreal2_cases[of m t])
+     (auto intro!: inj_onI simp: extreal_add_cancel_right extreal_mult_cancel_left)
+
+lemma extreal_PInfty_eq_plus[simp]:
+  shows "\<infinity> = a + b \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_MInfty_eq_plus[simp]:
+  shows "-\<infinity> = a + b \<longleftrightarrow> (a = -\<infinity> \<and> b \<noteq> \<infinity>) \<or> (b = -\<infinity> \<and> a \<noteq> \<infinity>)"
+  by (cases rule: extreal2_cases[of a b]) auto
+
+lemma extreal_less_divide_pos:
+  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y < z / x \<longleftrightarrow> x * y < z"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_divide_less_pos:
+  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y / x < z \<longleftrightarrow> y < x * z"
+  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma extreal_divide_eq:
+  "b \<noteq> 0 \<Longrightarrow> \<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> a / b = c \<longleftrightarrow> a = b * c"
+  by (cases rule: extreal3_cases[of a b c])
+     (simp_all add: field_simps)
+
+lemma extreal_inverse_not_MInfty[simp]: "inverse a \<noteq> -\<infinity>"
+  by (cases a) auto
+
+lemma extreal_mult_m1[simp]: "x * extreal (-1) = -x"
+  by (cases x) auto
+
+lemma extreal_LimI_finite:
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
+  assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r"
+  shows "u ----> x"
+proof (rule topological_tendstoI, unfold eventually_sequentially)
+  obtain rx where rx_def: "x=extreal rx" using assms by (cases x) auto
+  fix S assume "open S" "x : S"
+  then have "open (extreal -` S)" unfolding open_extreal_def by auto
+  with `x \<in> S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> extreal y \<in> S"
+    unfolding open_real_def rx_def by auto
+  then obtain n where
+    upper: "!!N. n <= N ==> u N < x + extreal r" and
+    lower: "!!N. n <= N ==> x < u N + extreal r" using assms(2)[of "extreal r"] by auto
+  show "EX N. ALL n>=N. u n : S"
+  proof (safe intro!: exI[of _ n])
+    fix N assume "n <= N"
+    from upper[OF this] lower[OF this] assms `0 < r`
+    have "u N ~: {\<infinity>,(-\<infinity>)}" by auto
+    from this obtain ra where ra_def: "(u N) = extreal ra" by (cases "u N") auto
+    hence "rx < ra + r" and "ra < rx + r"
+       using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto
+    hence "dist (real (u N)) rx < r"
+      using rx_def ra_def
+      by (auto simp: dist_real_def abs_diff_less_iff field_simps)
+    from dist[OF this] show "u N : S" using `u N  ~: {\<infinity>, -\<infinity>}`
+      by (auto simp: extreal_real split: split_if_asm)
+  qed
+qed
+
+lemma extreal_LimI_finite_iff:
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
+  shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))"
+  (is "?lhs <-> ?rhs")
+proof
+  assume lim: "u ----> x"
+  { fix r assume "(r::extreal)>0"
+    from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}"
+       apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
+       using lim extreal_between[of x r] assms `r>0` by auto
+    hence "EX N. ALL n>=N. u n < x + r & x < u n + r"
+      using extreal_minus_less[of r x] by (cases r) auto
+  } then show "?rhs" by auto
+next
+  assume ?rhs then show "u ----> x"
+    using extreal_LimI_finite[of x] assms by auto
+qed
+
+
+subsubsection {* @{text Liminf} and @{text Limsup} *}
+
+definition
+  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
+
+definition
+  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
+
+lemma Liminf_Sup:
+  fixes f :: "'a => 'b::{complete_lattice, linorder}"
+  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
+  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
+
+lemma Limsup_Inf:
+  fixes f :: "'a => 'b::{complete_lattice, linorder}"
+  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
+  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
+
+lemma extreal_SupI:
+  fixes x :: extreal
+  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
+  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
+  shows "Sup A = x"
+  unfolding Sup_extreal_def
+  using assms by (auto intro!: Least_equality)
+
+lemma extreal_InfI:
+  fixes x :: extreal
+  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
+  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
+  shows "Inf A = x"
+  unfolding Inf_extreal_def
+  using assms by (auto intro!: Greatest_equality)
+
+lemma Limsup_const:
+  fixes c :: "'a::{complete_lattice, linorder}"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Limsup net (\<lambda>x. c) = c"
+  unfolding Limsup_Inf
+proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
+  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
+  show "c \<le> x"
+  proof (rule ccontr)
+    assume "\<not> c \<le> x" then have "x < c" by auto
+    then show False using ntriv * by (auto simp: trivial_limit_def)
+  qed
+qed auto
+
+lemma Liminf_const:
+  fixes c :: "'a::{complete_lattice, linorder}"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Liminf net (\<lambda>x. c) = c"
+  unfolding Liminf_Sup
+proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
+  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
+  show "x \<le> c"
+  proof (rule ccontr)
+    assume "\<not> x \<le> c" then have "c < x" by auto
+    then show False using ntriv * by (auto simp: trivial_limit_def)
+  qed
+qed auto
+
+lemma mono_set:
+  fixes S :: "('a::order) set"
+  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
+  by (auto simp: mono_def mem_def)
+
+lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
+lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
+lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
+lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
+
+lemma mono_set_iff:
+  fixes S :: "'a::{linorder,complete_lattice} set"
+  defines "a \<equiv> Inf S"
+  shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
+proof
+  assume "mono S"
+  then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set)
+  show ?c
+  proof cases
+    assume "a \<in> S"
+    show ?c
+      using mono[OF _ `a \<in> S`]
+      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
+  next
+    assume "a \<notin> S"
+    have "S = {a <..}"
+    proof safe
+      fix x assume "x \<in> S"
+      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
+      then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
+    next
+      fix x assume "a < x"
+      then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff ..
+      with mono[of y x] show "x \<in> S" by auto
+    qed
+    then show ?c ..
+  qed
+qed auto
+
+lemma lim_imp_Liminf:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes ntriv: "\<not> trivial_limit net"
+  assumes lim: "(f ---> f0) net"
+  shows "Liminf net f = f0"
+  unfolding Liminf_Sup
+proof (safe intro!: extreal_SupI)
+  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
+  show "y \<le> f0"
+  proof (rule extreal_le_extreal)
+    fix B assume "B < y"
+    { assume "f0 < B"
+      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
+         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
+         by (auto intro: eventually_conj)
+      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
+      finally have False using ntriv[unfolded trivial_limit_def] by auto
+    } then show "B \<le> f0" by (metis linorder_le_less_linear)
+  qed
+next
+  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
+  show "f0 \<le> y"
+  proof (safe intro!: *[rule_format])
+    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
+      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
+  qed
+qed
+
+lemma extreal_Liminf_le_Limsup:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Liminf net f \<le> Limsup net f"
+  unfolding Limsup_Inf Liminf_Sup
+proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
+  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
+  show "u \<le> v"
+  proof (rule ccontr)
+    assume "\<not> u \<le> v"
+    then obtain t where "t < u" "v < t"
+      using extreal_dense[of v u] by (auto simp: not_le)
+    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
+      using * by (auto intro: eventually_conj)
+    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
+    finally show False using ntriv by (auto simp: trivial_limit_def)
+  qed
+qed
+
+lemma Liminf_mono:
+  fixes f g :: "'a => extreal"
+  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
+  shows "Liminf net f \<le> Liminf net g"
+  unfolding Liminf_Sup
+proof (safe intro!: Sup_mono bexI)
+  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
+  then have "eventually (\<lambda>x. y < f x) net" by auto
+  then show "eventually (\<lambda>x. y < g x) net"
+    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
+qed simp
+
+lemma Liminf_eq:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  assumes "eventually (\<lambda>x. f x = g x) net"
+  shows "Liminf net f = Liminf net g"
+  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
+
+lemma Liminf_mono_all:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  assumes "\<And>x. f x \<le> g x"
+  shows "Liminf net f \<le> Liminf net g"
+  using assms by (intro Liminf_mono always_eventually) auto
+
+lemma Limsup_mono:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
+  shows "Limsup net f \<le> Limsup net g"
+  unfolding Limsup_Inf
+proof (safe intro!: Inf_mono bexI)
+  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
+  then have "eventually (\<lambda>x. g x < y) net" by auto
+  then show "eventually (\<lambda>x. f x < y) net"
+    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
+qed simp
+
+lemma Limsup_mono_all:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  assumes "\<And>x. f x \<le> g x"
+  shows "Limsup net f \<le> Limsup net g"
+  using assms by (intro Limsup_mono always_eventually) auto
+
+lemma Limsup_eq:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  assumes "eventually (\<lambda>x. f x = g x) net"
+  shows "Limsup net f = Limsup net g"
+  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
+
+abbreviation "liminf \<equiv> Liminf sequentially"
+
+abbreviation "limsup \<equiv> Limsup sequentially"
+
+lemma (in complete_lattice) less_INFD:
+  assumes "y < INFI A f"" i \<in> A" shows "y < f i"
+proof -
+  note `y < INFI A f`
+  also have "INFI A f \<le> f i" using `i \<in> A` by (rule INF_leI)
+  finally show "y < f i" .
+qed
+
+lemma liminf_SUPR_INFI:
+  fixes f :: "nat \<Rightarrow> extreal"
+  shows "liminf f = (SUP n. INF m:{n..}. f m)"
+  unfolding Liminf_Sup eventually_sequentially
+proof (safe intro!: antisym complete_lattice_class.Sup_least)
+  fix x assume *: "\<forall>y<x. \<exists>N. \<forall>n\<ge>N. y < f n" show "x \<le> (SUP n. INF m:{n..}. f m)"
+  proof (rule extreal_le_extreal)
+    fix y assume "y < x"
+    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
+    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
+    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto
+    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
+  qed
+next
+  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
+  proof (unfold SUPR_def, safe intro!: Sup_mono bexI)
+    fix y n assume "y < INFI {n..} f"
+    from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
+  qed (rule order_refl)
+qed
+
+lemma tail_same_limsup:
+  fixes X Y :: "nat => extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
+  shows "limsup X = limsup Y"
+  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma tail_same_liminf:
+  fixes X Y :: "nat => extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
+  shows "liminf X = liminf Y"
+  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma liminf_mono:
+  fixes X Y :: "nat \<Rightarrow> extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
+  shows "liminf X \<le> liminf Y"
+  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma limsup_mono:
+  fixes X Y :: "nat => extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
+  shows "limsup X \<le> limsup Y"
+  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
+
+declare trivial_limit_sequentially[simp]
+
+lemma
+  fixes X :: "nat \<Rightarrow> extreal"
+  shows extreal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
+    and extreal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
+  unfolding incseq_def decseq_def by auto
+
+lemma liminf_bounded:
+  fixes X Y :: "nat \<Rightarrow> extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
+  shows "C \<le> liminf X"
+  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
+
+lemma limsup_bounded:
+  fixes X Y :: "nat => extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
+  shows "limsup X \<le> C"
+  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
+
+lemma liminf_bounded_iff:
+  fixes x :: "nat \<Rightarrow> extreal"
+  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
+proof safe
+  fix B assume "B < C" "C \<le> liminf x"
+  then have "B < liminf x" by auto
+  then obtain N where "B < (INF m:{N..}. x m)"
+    unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto
+  from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
+next
+  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
+  { fix B assume "B<C"
+    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
+    hence "B \<le> (INF m:{N..}. x m)" by (intro le_INFI) auto
+    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp
+    finally have "B \<le> liminf x" .
+  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
+qed
+
+lemma liminf_subseq_mono:
+  fixes X :: "nat \<Rightarrow> extreal"
+  assumes "subseq r"
+  shows "liminf X \<le> liminf (X \<circ> r) "
+proof-
+  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
+  proof (safe intro!: INF_mono)
+    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
+      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
+  qed
+  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
+qed
+
+lemma extreal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "extreal (real x) = x"
+  using assms by auto
+
+lemma extreal_le_extreal_bounded:
+  fixes x y z :: extreal
+  assumes "z \<le> y"
+  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
+  shows "x \<le> y"
+proof (rule extreal_le_extreal)
+  fix B assume "B < x"
+  show "B \<le> y"
+  proof cases
+    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
+  next
+    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
+  qed
+qed
+
+lemma fixes x y :: extreal
+  shows Sup_atMost[simp]: "Sup {.. y} = y"
+    and Sup_lessThan[simp]: "Sup {..< y} = y"
+    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
+    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
+    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
+  by (auto simp: Sup_extreal_def intro!: Least_equality
+           intro: extreal_le_extreal extreal_le_extreal_bounded[of x])
+
+lemma Sup_greaterThanlessThan[simp]:
+  fixes x y :: extreal assumes "x < y" shows "Sup { x <..< y} = y"
+  unfolding Sup_extreal_def
+proof (intro Least_equality extreal_le_extreal_bounded[of _ _ y])
+  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
+  from extreal_dense[OF `x < y`] guess w .. note w = this
+  with z[THEN bspec, of w] show "x \<le> z" by auto
+qed auto
+
+lemma real_extreal_id: "real o extreal = id"
+proof-
+{ fix x have "(real o extreal) x = id x" by auto }
+from this show ?thesis using ext by blast
+qed
+
+
+lemma open_image_extreal: "open(UNIV-{\<infinity>,(-\<infinity>)})"
+by (metis range_extreal open_extreal open_UNIV)
+
+lemma extreal_le_distrib:
+  fixes a b c :: extreal shows "c * (a + b) \<le> c * a + c * b"
+  by (cases rule: extreal3_cases[of a b c])
+     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
+
+lemma extreal_pos_distrib:
+  fixes a b c :: extreal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
+  using assms by (cases rule: extreal3_cases[of a b c])
+                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
+
+lemma extreal_pos_le_distrib:
+fixes a b c :: extreal
+assumes "c>=0"
+shows "c * (a + b) <= c * a + c * b"
+  using assms by (cases rule: extreal3_cases[of a b c])
+                 (auto simp add: field_simps)
+
+lemma extreal_max_mono:
+  "[| (a::extreal) <= b; c <= d |] ==> max a c <= max b d"
+  by (metis sup_extreal_def sup_mono)
+
+
+lemma extreal_max_least:
+  "[| (a::extreal) <= x; c <= x |] ==> max a c <= x"
+  by (metis sup_extreal_def sup_least)
+
+end
--- a/src/HOL/Limits.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Limits.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -103,7 +103,6 @@
   shows "eventually (\<lambda>i. R i) net"
 using assms by (auto elim!: eventually_rev_mp)
 
-
 subsection {* Finer-than relation *}
 
 text {* @{term "net \<le> net'"} means that @{term net} is finer than
@@ -231,7 +230,6 @@
   "eventually (\<lambda>x. False) net \<longleftrightarrow> net = bot"
 unfolding expand_net_eq by (auto elim: eventually_rev_mp)
 
-
 subsection {* Map function for nets *}
 
 definition netmap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a net \<Rightarrow> 'b net" where
@@ -287,6 +285,13 @@
 by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp)
 
 
+definition
+  trivial_limit :: "'a net \<Rightarrow> bool" where
+  "trivial_limit net \<longleftrightarrow> eventually (\<lambda>x. False) net"
+
+lemma trivial_limit_sequentially[intro]: "\<not> trivial_limit sequentially"
+  by (auto simp add: trivial_limit_def eventually_sequentially)
+
 subsection {* Standard Nets *}
 
 definition within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70) where
@@ -827,4 +832,29 @@
     \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) net"
 by (simp add: mult.tendsto tendsto_inverse divide_inverse)
 
+lemma tendsto_unique:
+  fixes f :: "'a \<Rightarrow> 'b::t2_space"
+  assumes "\<not> trivial_limit net"  "(f ---> l) net"  "(f ---> l') net"
+  shows "l = l'"
+proof (rule ccontr)
+  assume "l \<noteq> l'"
+  obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
+    using hausdorff [OF `l \<noteq> l'`] by fast
+  have "eventually (\<lambda>x. f x \<in> U) net"
+    using `(f ---> l) net` `open U` `l \<in> U` by (rule topological_tendstoD)
+  moreover
+  have "eventually (\<lambda>x. f x \<in> V) net"
+    using `(f ---> l') net` `open V` `l' \<in> V` by (rule topological_tendstoD)
+  ultimately
+  have "eventually (\<lambda>x. False) net"
+  proof (rule eventually_elim2)
+    fix x
+    assume "f x \<in> U" "f x \<in> V"
+    hence "f x \<in> U \<inter> V" by simp
+    with `U \<inter> V = {}` show "False" by simp
+  qed
+  with `\<not> trivial_limit net` show "False"
+    by (simp add: trivial_limit_def)
+qed
+
 end
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1129,11 +1129,11 @@
     show "bounded_linear (g' x)" unfolding linear_linear linear_def apply(rule,rule,rule) defer proof(rule,rule)
       fix x' y z::"'m" and c::real
       note lin = assms(2)[rule_format,OF `x\<in>s`,THEN derivative_linear]
-      show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply(rule Lim_unique[OF trivial_limit_sequentially])
+      show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply(rule tendsto_unique[OF trivial_limit_sequentially])
         apply(rule lem3[rule_format])
         unfolding lin[unfolded bounded_linear_def bounded_linear_axioms_def,THEN conjunct2,THEN conjunct1,rule_format]
         apply(rule Lim_cmul) by(rule lem3[rule_format])
-      show "g' x (y + z) = g' x y + g' x z" apply(rule Lim_unique[OF trivial_limit_sequentially])
+      show "g' x (y + z) = g' x y + g' x z" apply(rule tendsto_unique[OF trivial_limit_sequentially])
         apply(rule lem3[rule_format]) unfolding lin[unfolded bounded_linear_def additive_def,THEN conjunct1,rule_format]
         apply(rule Lim_add) by(rule lem3[rule_format])+ qed 
     show "\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm (y - x) < d \<longrightarrow> norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" proof(rule,rule) case goal1
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -326,42 +326,6 @@
 
 text{* Hence more metric properties. *}
 
-lemma dist_triangle_alt:
-  fixes x y z :: "'a::metric_space"
-  shows "dist y z <= dist x y + dist x z"
-by (rule dist_triangle3)
-
-lemma dist_pos_lt:
-  fixes x y :: "'a::metric_space"
-  shows "x \<noteq> y ==> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_nz:
-  fixes x y :: "'a::metric_space"
-  shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
-by (simp add: zero_less_dist_iff)
-
-lemma dist_triangle_le:
-  fixes x y z :: "'a::metric_space"
-  shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
-by (rule order_trans [OF dist_triangle2])
-
-lemma dist_triangle_lt:
-  fixes x y z :: "'a::metric_space"
-  shows "dist x z + dist y z < e ==> dist x y < e"
-by (rule le_less_trans [OF dist_triangle2])
-
-lemma dist_triangle_half_l:
-  fixes x1 x2 y :: "'a::metric_space"
-  shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_lt [where z=y], simp)
-
-lemma dist_triangle_half_r:
-  fixes x1 x2 y :: "'a::metric_space"
-  shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
-by (rule dist_triangle_half_l, simp_all add: dist_commute)
-
-
 lemma norm_triangle_half_r:
   shows "norm (y - x1) < e / 2 \<Longrightarrow> norm (y - x2) < e / 2 \<Longrightarrow> norm (x1 - x2) < e"
   using dist_triangle_half_r unfolding dist_norm[THEN sym] by auto
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -0,0 +1,1261 @@
+(* Title: src/HOL/Multivariate_Analysis/Extended_Reals.thy
+   Author: Johannes Hölzl; TU München
+   Author: Robert Himmelmann; TU München
+   Author: Armin Heller; TU München
+   Author: Bogdan Grechuk; University of Edinburgh *)
+
+header {* Limits on the Extended real number line *}
+
+theory Extended_Real_Limits
+  imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Reals"
+begin
+
+lemma continuous_on_extreal[intro, simp]: "continuous_on A extreal"
+  unfolding continuous_on_topological open_extreal_def by auto
+
+lemma continuous_at_extreal[intro, simp]: "continuous (at x) extreal"
+  using continuous_on_eq_continuous_at[of UNIV] by auto
+
+lemma continuous_within_extreal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) extreal"
+  using continuous_on_eq_continuous_within[of A] by auto
+
+lemma extreal_open_uminus:
+  fixes S :: "extreal set"
+  assumes "open S"
+  shows "open (uminus ` S)"
+  unfolding open_extreal_def
+proof (intro conjI impI)
+  obtain x y where S: "open (extreal -` S)"
+    "\<infinity> \<in> S \<Longrightarrow> {extreal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< extreal y} \<subseteq> S"
+    using `open S` unfolding open_extreal_def by auto
+  have "extreal -` uminus ` S = uminus ` (extreal -` S)"
+  proof safe
+    fix x y assume "extreal x = - y" "y \<in> S"
+    then show "x \<in> uminus ` extreal -` S" by (cases y) auto
+  next
+    fix x assume "extreal x \<in> S"
+    then show "- x \<in> extreal -` uminus ` S"
+      by (auto intro: image_eqI[of _ _ "extreal x"])
+  qed
+  then show "open (extreal -` uminus ` S)"
+    using S by (auto intro: open_negations)
+  { assume "\<infinity> \<in> uminus ` S"
+    then have "-\<infinity> \<in> S" by (metis image_iff extreal_uminus_uminus)
+    then have "uminus ` {..<extreal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
+    then show "\<exists>x. {extreal x<..} \<subseteq> uminus ` S" using extreal_uminus_lessThan by auto }
+  { assume "-\<infinity> \<in> uminus ` S"
+    then have "\<infinity> : S" by (metis image_iff extreal_uminus_uminus)
+    then have "uminus ` {extreal x<..} <= uminus ` S" using S by (intro image_mono) auto
+    then show "\<exists>y. {..<extreal y} <= uminus ` S" using extreal_uminus_greaterThan by auto }
+qed
+
+lemma extreal_uminus_complement:
+  fixes S :: "extreal set"
+  shows "uminus ` (- S) = - uminus ` S"
+  by (auto intro!: bij_image_Compl_eq surjI[of _ uminus] simp: bij_betw_def)
+
+lemma extreal_closed_uminus:
+  fixes S :: "extreal set"
+  assumes "closed S"
+  shows "closed (uminus ` S)"
+using assms unfolding closed_def
+using extreal_open_uminus[of "- S"] extreal_uminus_complement by auto
+
+lemma not_open_extreal_singleton:
+  "\<not> (open {a :: extreal})"
+proof(rule ccontr)
+  assume "\<not> \<not> open {a}" hence a: "open {a}" by auto
+  show False
+  proof (cases a)
+    case MInf
+    then obtain y where "{..<extreal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
+    hence "extreal(y - 1):{a}" apply (subst subsetD[of "{..<extreal y}"]) by auto
+    then show False using `a=(-\<infinity>)` by auto
+  next
+    case PInf
+    then obtain y where "{extreal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
+    hence "extreal(y+1):{a}" apply (subst subsetD[of "{extreal y<..}"]) by auto
+    then show False using `a=\<infinity>` by auto
+  next
+    case (real r) then have fin: "\<bar>a\<bar> \<noteq> \<infinity>" by simp
+    from extreal_open_cont_interval[OF a singletonI this] guess e . note e = this
+    then obtain b where b_def: "a<b & b<a+e"
+      using fin extreal_between extreal_dense[of a "a+e"] by auto
+    then have "b: {a-e <..< a+e}" using fin extreal_between[of a e] e by auto
+    then show False using b_def e by auto
+  qed
+qed
+
+lemma extreal_closed_contains_Inf:
+  fixes S :: "extreal set"
+  assumes "closed S" "S ~= {}"
+  shows "Inf S : S"
+proof(rule ccontr)
+  assume "Inf S \<notin> S" hence a: "open (-S)" "Inf S:(- S)" using assms by auto
+  show False
+  proof (cases "Inf S")
+    case MInf hence "(-\<infinity>) : - S" using a by auto
+    then obtain y where "{..<extreal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
+    hence "extreal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff
+      complete_lattice_class.Inf_greatest double_complement set_rev_mp)
+    then show False using MInf by auto
+  next
+    case PInf then have "S={\<infinity>}" by (metis Inf_eq_PInfty assms(2))
+    then show False by (metis `Inf S ~: S` insert_code mem_def PInf)
+  next
+    case (real r) then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp
+    from extreal_open_cont_interval[OF a this] guess e . note e = this
+    { fix x assume "x:S" hence "x>=Inf S" by (rule complete_lattice_class.Inf_lower)
+      hence *: "x>Inf S-e" using e by (metis fin extreal_between(1) order_less_le_trans)
+      { assume "x<Inf S+e" hence "x:{Inf S-e <..< Inf S+e}" using * by auto
+        hence False using e `x:S` by auto
+      } hence "x>=Inf S+e" by (metis linorder_le_less_linear)
+    } hence "Inf S + e <= Inf S" by (metis le_Inf_iff)
+    then show False using real e by (cases e) auto
+  qed
+qed
+
+lemma extreal_closed_contains_Sup:
+  fixes S :: "extreal set"
+  assumes "closed S" "S ~= {}"
+  shows "Sup S : S"
+proof-
+  have "closed (uminus ` S)" by (metis assms(1) extreal_closed_uminus)
+  hence "Inf (uminus ` S) : uminus ` S" using assms extreal_closed_contains_Inf[of "uminus ` S"] by auto
+  hence "- Sup S : uminus ` S" using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image)
+  thus ?thesis by (metis imageI extreal_uminus_uminus extreal_minus_minus_image)
+qed
+
+lemma extreal_open_closed_aux:
+  fixes S :: "extreal set"
+  assumes "open S" "closed S"
+  assumes S: "(-\<infinity>) ~: S"
+  shows "S = {}"
+proof(rule ccontr)
+  assume "S ~= {}"
+  hence *: "(Inf S):S" by (metis assms(2) extreal_closed_contains_Inf)
+  { assume "Inf S=(-\<infinity>)" hence False using * assms(3) by auto }
+  moreover
+  { assume "Inf S=\<infinity>" hence "S={\<infinity>}" by (metis Inf_eq_PInfty `S ~= {}`)
+    hence False by (metis assms(1) not_open_extreal_singleton) }
+  moreover
+  { assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"
+    from extreal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
+    then obtain b where b_def: "Inf S-e<b & b<Inf S"
+      using fin extreal_between[of "Inf S" e] extreal_dense[of "Inf S-e"] by auto
+    hence "b: {Inf S-e <..< Inf S+e}" using e fin extreal_between[of "Inf S" e] by auto
+    hence "b:S" using e by auto
+    hence False using b_def by (metis complete_lattice_class.Inf_lower leD)
+  } ultimately show False by auto
+qed
+
+lemma extreal_open_closed:
+  fixes S :: "extreal set"
+  shows "(open S & closed S) <-> (S = {} | S = UNIV)"
+proof-
+{ assume lhs: "open S & closed S"
+  { assume "(-\<infinity>) ~: S" hence "S={}" using lhs extreal_open_closed_aux by auto }
+  moreover
+  { assume "(-\<infinity>) : S" hence "(- S)={}" using lhs extreal_open_closed_aux[of "-S"] by auto }
+  ultimately have "S = {} | S = UNIV" by auto
+} thus ?thesis by auto
+qed
+
+lemma extreal_open_affinity_pos:
+  assumes "open S" and m: "m \<noteq> \<infinity>" "0 < m" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
+  shows "open ((\<lambda>x. m * x + t) ` S)"
+proof -
+  obtain r where r[simp]: "m = extreal r" using m by (cases m) auto
+  obtain p where p[simp]: "t = extreal p" using t by auto
+  have "r \<noteq> 0" "0 < r" and m': "m \<noteq> \<infinity>" "m \<noteq> -\<infinity>" "m \<noteq> 0" using m by auto
+  from `open S`[THEN extreal_openE] guess l u . note T = this
+  let ?f = "(\<lambda>x. m * x + t)"
+  show ?thesis unfolding open_extreal_def
+  proof (intro conjI impI exI subsetI)
+    have "extreal -` ?f ` S = (\<lambda>x. r * x + p) ` (extreal -` S)"
+    proof safe
+      fix x y assume "extreal y = m * x + t" "x \<in> S"
+      then show "y \<in> (\<lambda>x. r * x + p) ` extreal -` S"
+        using `r \<noteq> 0` by (cases x) (auto intro!: image_eqI[of _ _ "real x"] split: split_if_asm)
+    qed force
+    then show "open (extreal -` ?f ` S)"
+      using open_affinity[OF T(1) `r \<noteq> 0`] by (auto simp: ac_simps)
+  next
+    assume "\<infinity> \<in> ?f`S" with `0 < r` have "\<infinity> \<in> S" by auto
+    fix x assume "x \<in> {extreal (r * l + p)<..}"
+    then have [simp]: "extreal (r * l + p) < x" by auto
+    show "x \<in> ?f`S"
+    proof (rule image_eqI)
+      show "x = m * ((x - t) / m) + t"
+        using m t by (cases rule: extreal3_cases[of m x t]) auto
+      have "extreal l < (x - t)/m"
+        using m t by (simp add: extreal_less_divide_pos extreal_less_minus)
+      then show "(x - t)/m \<in> S" using T(2)[OF `\<infinity> \<in> S`] by auto
+    qed
+  next
+    assume "-\<infinity> \<in> ?f`S" with `0 < r` have "-\<infinity> \<in> S" by auto
+    fix x assume "x \<in> {..<extreal (r * u + p)}"
+    then have [simp]: "x < extreal (r * u + p)" by auto
+    show "x \<in> ?f`S"
+    proof (rule image_eqI)
+      show "x = m * ((x - t) / m) + t"
+        using m t by (cases rule: extreal3_cases[of m x t]) auto
+      have "(x - t)/m < extreal u"
+        using m t by (simp add: extreal_divide_less_pos extreal_minus_less)
+      then show "(x - t)/m \<in> S" using T(3)[OF `-\<infinity> \<in> S`] by auto
+    qed
+  qed
+qed
+
+lemma extreal_open_affinity:
+  assumes "open S" and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
+  shows "open ((\<lambda>x. m * x + t) ` S)"
+proof cases
+  assume "0 < m" then show ?thesis
+    using extreal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto
+next
+  assume "\<not> 0 < m" then
+  have "0 < -m" using `m \<noteq> 0` by (cases m) auto
+  then have m: "-m \<noteq> \<infinity>" "0 < -m" using `\<bar>m\<bar> \<noteq> \<infinity>`
+    by (auto simp: extreal_uminus_eq_reorder)
+  from extreal_open_affinity_pos[OF extreal_open_uminus[OF `open S`] m t]
+  show ?thesis unfolding image_image by simp
+qed
+
+lemma extreal_lim_mult:
+  fixes X :: "'a \<Rightarrow> extreal"
+  assumes lim: "(X ---> L) net" and a: "\<bar>a\<bar> \<noteq> \<infinity>"
+  shows "((\<lambda>i. a * X i) ---> a * L) net"
+proof cases
+  assume "a \<noteq> 0"
+  show ?thesis
+  proof (rule topological_tendstoI)
+    fix S assume "open S" "a * L \<in> S"
+    have "a * L / a = L"
+      using `a \<noteq> 0` a by (cases rule: extreal2_cases[of a L]) auto
+    then have L: "L \<in> ((\<lambda>x. x / a) ` S)"
+      using `a * L \<in> S` by (force simp: image_iff)
+    moreover have "open ((\<lambda>x. x / a) ` S)"
+      using extreal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a
+      by (auto simp: extreal_divide_eq extreal_inverse_eq_0 divide_extreal_def ac_simps)
+    note * = lim[THEN topological_tendstoD, OF this L]
+    { fix x from a `a \<noteq> 0` have "a * (x / a) = x"
+        by (cases rule: extreal2_cases[of a x]) auto }
+    note this[simp]
+    show "eventually (\<lambda>x. a * X x \<in> S) net"
+      by (rule eventually_mono[OF _ *]) auto
+  qed
+qed auto
+
+lemma extreal_lim_uminus:
+  fixes X :: "'a \<Rightarrow> extreal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
+  using extreal_lim_mult[of X L net "extreal (-1)"]
+        extreal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "extreal (-1)"]
+  by (auto simp add: algebra_simps)
+
+lemma Lim_bounded2_extreal:
+  assumes lim:"f ----> (l :: extreal)"
+  and ge: "ALL n>=N. f n >= C"
+  shows "l>=C"
+proof-
+def g == "(%i. -(f i))"
+{ fix n assume "n>=N" hence "g n <= -C" using assms extreal_minus_le_minus g_def by auto }
+hence "ALL n>=N. g n <= -C" by auto
+moreover have limg: "g ----> (-l)" using g_def extreal_lim_uminus lim by auto
+ultimately have "-l <= -C" using Lim_bounded_extreal[of g "-l" _ "-C"] by auto
+from this show ?thesis using extreal_minus_le_minus by auto
+qed
+
+
+lemma extreal_open_atLeast: "open {x..} \<longleftrightarrow> x = -\<infinity>"
+proof
+  assume "x = -\<infinity>" then have "{x..} = UNIV" by auto
+  then show "open {x..}" by auto
+next
+  assume "open {x..}"
+  then have "open {x..} \<and> closed {x..}" by auto
+  then have "{x..} = UNIV" unfolding extreal_open_closed by auto
+  then show "x = -\<infinity>" by (simp add: bot_extreal_def atLeast_eq_UNIV_iff)
+qed
+
+lemma extreal_open_mono_set:
+  fixes S :: "extreal set"
+  defines "a \<equiv> Inf S"
+  shows "(open S \<and> mono S) \<longleftrightarrow> (S = UNIV \<or> S = {a <..})"
+  by (metis Inf_UNIV a_def atLeast_eq_UNIV_iff extreal_open_atLeast
+            extreal_open_closed mono_set_iff open_extreal_greaterThan)
+
+lemma extreal_closed_mono_set:
+  fixes S :: "extreal set"
+  shows "(closed S \<and> mono S) \<longleftrightarrow> (S = {} \<or> S = {Inf S ..})"
+  by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_extreal_atLeast
+            extreal_open_closed mono_empty mono_set_iff open_extreal_greaterThan)
+
+lemma extreal_Liminf_Sup_monoset:
+  fixes f :: "'a => extreal"
+  shows "Liminf net f = Sup {l. \<forall>S. open S \<longrightarrow> mono S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
+  unfolding Liminf_Sup
+proof (intro arg_cong[where f="\<lambda>P. Sup (Collect P)"] ext iffI allI impI)
+  fix l S assume ev: "\<forall>y<l. eventually (\<lambda>x. y < f x) net" and "open S" "mono S" "l \<in> S"
+  then have "S = UNIV \<or> S = {Inf S <..}"
+    using extreal_open_mono_set[of S] by auto
+  then show "eventually (\<lambda>x. f x \<in> S) net"
+  proof
+    assume S: "S = {Inf S<..}"
+    then have "Inf S < l" using `l \<in> S` by auto
+    then have "eventually (\<lambda>x. Inf S < f x) net" using ev by auto
+    then show "eventually (\<lambda>x. f x \<in> S) net"  by (subst S) auto
+  qed auto
+next
+  fix l y assume S: "\<forall>S. open S \<longrightarrow> mono S \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net" "y < l"
+  have "eventually  (\<lambda>x. f x \<in> {y <..}) net"
+    using `y < l` by (intro S[rule_format]) auto
+  then show "eventually (\<lambda>x. y < f x) net" by auto
+qed
+
+lemma extreal_Limsup_Inf_monoset:
+  fixes f :: "'a => extreal"
+  shows "Limsup net f = Inf {l. \<forall>S. open S \<longrightarrow> mono (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
+  unfolding Limsup_Inf
+proof (intro arg_cong[where f="\<lambda>P. Inf (Collect P)"] ext iffI allI impI)
+  fix l S assume ev: "\<forall>y>l. eventually (\<lambda>x. f x < y) net" and "open S" "mono (uminus`S)" "l \<in> S"
+  then have "open (uminus`S) \<and> mono (uminus`S)" by (simp add: extreal_open_uminus)
+  then have "S = UNIV \<or> S = {..< Sup S}"
+    unfolding extreal_open_mono_set extreal_Inf_uminus_image_eq extreal_image_uminus_shift by simp
+  then show "eventually (\<lambda>x. f x \<in> S) net"
+  proof
+    assume S: "S = {..< Sup S}"
+    then have "l < Sup S" using `l \<in> S` by auto
+    then have "eventually (\<lambda>x. f x < Sup S) net" using ev by auto
+    then show "eventually (\<lambda>x. f x \<in> S) net"  by (subst S) auto
+  qed auto
+next
+  fix l y assume S: "\<forall>S. open S \<longrightarrow> mono (uminus`S) \<longrightarrow> l \<in> S \<longrightarrow> eventually  (\<lambda>x. f x \<in> S) net" "l < y"
+  have "eventually  (\<lambda>x. f x \<in> {..< y}) net"
+    using `l < y` by (intro S[rule_format]) auto
+  then show "eventually (\<lambda>x. f x < y) net" by auto
+qed
+
+
+lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::extreal set)"
+  using extreal_open_uminus[of S] extreal_open_uminus[of "uminus`S"] by auto
+
+lemma extreal_Limsup_uminus:
+  fixes f :: "'a => extreal"
+  shows "Limsup net (\<lambda>x. - (f x)) = -(Liminf net f)"
+proof -
+  { fix P l have "(\<exists>x. (l::extreal) = -x \<and> P x) \<longleftrightarrow> P (-l)" by (auto intro!: exI[of _ "-l"]) }
+  note Ex_cancel = this
+  { fix P :: "extreal set \<Rightarrow> bool" have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))"
+      apply auto by (erule_tac x="uminus`S" in allE) (auto simp: image_image) }
+  note add_uminus_image = this
+  { fix x S have "(x::extreal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" by (auto intro!: image_eqI[of _ _ "-x"]) }
+  note remove_uminus_image = this
+  show ?thesis
+    unfolding extreal_Limsup_Inf_monoset extreal_Liminf_Sup_monoset
+    unfolding extreal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel
+    by (subst add_uminus_image) (simp add: open_uminus_iff remove_uminus_image)
+qed
+
+lemma extreal_Liminf_uminus:
+  fixes f :: "'a => extreal"
+  shows "Liminf net (\<lambda>x. - (f x)) = -(Limsup net f)"
+  using extreal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto
+
+lemma extreal_Lim_uminus:
+  fixes f :: "'a \<Rightarrow> extreal" shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net"
+  using
+    extreal_lim_mult[of f f0 net "- 1"]
+    extreal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"]
+  by (auto simp: extreal_uminus_reorder)
+
+lemma lim_imp_Limsup:
+  fixes f :: "'a => extreal"
+  assumes "\<not> trivial_limit net"
+  assumes lim: "(f ---> f0) net"
+  shows "Limsup net f = f0"
+  using extreal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"]
+     extreal_Liminf_uminus[of net f] assms by simp
+
+lemma Liminf_PInfty:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes "\<not> trivial_limit net"
+  shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>"
+proof (intro lim_imp_Liminf iffI assms)
+  assume rhs: "Liminf net f = \<infinity>"
+  { fix S assume "open S & \<infinity> : S"
+    then obtain m where "{extreal m<..} <= S" using open_PInfty2 by auto
+    moreover have "eventually (\<lambda>x. f x \<in> {extreal m<..}) net"
+      using rhs unfolding Liminf_Sup top_extreal_def[symmetric] Sup_eq_top_iff
+      by (auto elim!: allE[where x="extreal m"] simp: top_extreal_def)
+    ultimately have "eventually (%x. f x : S) net" apply (subst eventually_mono) by auto
+  } then show "(f ---> \<infinity>) net" unfolding tendsto_def by auto
+qed
+
+lemma Limsup_MInfty:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes "\<not> trivial_limit net"
+  shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>"
+  using assms extreal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"]
+        extreal_Liminf_uminus[of _ f] by (auto simp: extreal_uminus_eq_reorder)
+
+lemma extreal_Liminf_eq_Limsup:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes ntriv: "\<not> trivial_limit net"
+  assumes lim: "Liminf net f = f0" "Limsup net f = f0"
+  shows "(f ---> f0) net"
+proof (cases f0)
+  case PInf then show ?thesis using Liminf_PInfty[OF ntriv] lim by auto
+next
+  case MInf then show ?thesis using Limsup_MInfty[OF ntriv] lim by auto
+next
+  case (real r)
+  show "(f ---> f0) net"
+  proof (rule topological_tendstoI)
+    fix S assume "open S""f0 \<in> S"
+    then obtain a b where "a < Liminf net f" "Limsup net f < b" "{a<..<b} \<subseteq> S"
+      using extreal_open_cont_interval2[of S f0] real lim by auto
+    then have "eventually (\<lambda>x. f x \<in> {a<..<b}) net"
+      unfolding Liminf_Sup Limsup_Inf less_Sup_iff Inf_less_iff
+      by (auto intro!: eventually_conj simp add: greaterThanLessThan_iff)
+    with `{a<..<b} \<subseteq> S` show "eventually (%x. f x : S) net"
+      by (rule_tac eventually_mono) auto
+  qed
+qed
+
+lemma extreal_Liminf_eq_Limsup_iff:
+  fixes f :: "'a \<Rightarrow> extreal"
+  assumes "\<not> trivial_limit net"
+  shows "(f ---> f0) net \<longleftrightarrow> Liminf net f = f0 \<and> Limsup net f = f0"
+  by (metis assms extreal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup)
+
+lemma limsup_INFI_SUPR:
+  fixes f :: "nat \<Rightarrow> extreal"
+  shows "limsup f = (INF n. SUP m:{n..}. f m)"
+  using extreal_Limsup_uminus[of sequentially "\<lambda>x. - f x"]
+  by (simp add: liminf_SUPR_INFI extreal_INFI_uminus extreal_SUPR_uminus)
+
+lemma liminf_PInfty:
+  fixes X :: "nat => extreal"
+  shows "X ----> \<infinity> <-> liminf X = \<infinity>"
+by (metis Liminf_PInfty trivial_limit_sequentially)
+
+lemma limsup_MInfty:
+  fixes X :: "nat => extreal"
+  shows "X ----> (-\<infinity>) <-> limsup X = (-\<infinity>)"
+by (metis Limsup_MInfty trivial_limit_sequentially)
+
+lemma extreal_lim_mono:
+  fixes X Y :: "nat => extreal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
+  assumes "X ----> x" "Y ----> y"
+  shows "x <= y"
+  by (metis extreal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
+
+lemma incseq_le_extreal:
+  fixes X :: "nat \<Rightarrow> extreal"
+  assumes inc: "incseq X" and lim: "X ----> L"
+  shows "X N \<le> L"
+  using inc
+  by (intro extreal_lim_mono[of N, OF _ Lim_const lim]) (simp add: incseq_def)
+
+lemma decseq_ge_extreal: assumes dec: "decseq X"
+  and lim: "X ----> (L::extreal)" shows "X N >= L"
+  using dec
+  by (intro extreal_lim_mono[of N, OF _ lim Lim_const]) (simp add: decseq_def)
+
+lemma liminf_bounded_open:
+  fixes x :: "nat \<Rightarrow> extreal"
+  shows "x0 \<le> liminf x \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> mono S \<longrightarrow> x0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. x n \<in> S))" 
+  (is "_ \<longleftrightarrow> ?P x0")
+proof
+  assume "?P x0" then show "x0 \<le> liminf x"
+    unfolding extreal_Liminf_Sup_monoset eventually_sequentially
+    by (intro complete_lattice_class.Sup_upper) auto
+next
+  assume "x0 \<le> liminf x"
+  { fix S :: "extreal set" assume om: "open S & mono S & x0:S"
+    { assume "S = UNIV" hence "EX N. (ALL n>=N. x n : S)" by auto }
+    moreover
+    { assume "~(S=UNIV)"
+      then obtain B where B_def: "S = {B<..}" using om extreal_open_mono_set by auto
+      hence "B<x0" using om by auto
+      hence "EX N. ALL n>=N. x n : S" unfolding B_def using `x0 \<le> liminf x` liminf_bounded_iff by auto
+    } ultimately have "EX N. (ALL n>=N. x n : S)" by auto
+  } then show "?P x0" by auto
+qed
+
+lemma limsup_subseq_mono:
+  fixes X :: "nat \<Rightarrow> extreal"
+  assumes "subseq r"
+  shows "limsup (X \<circ> r) \<le> limsup X"
+proof-
+  have "(\<lambda>n. - X n) \<circ> r = (\<lambda>n. - (X \<circ> r) n)" by (simp add: fun_eq_iff)
+  then have "- limsup X \<le> - limsup (X \<circ> r)"
+     using liminf_subseq_mono[of r "(%n. - X n)"]
+       extreal_Liminf_uminus[of sequentially X]
+       extreal_Liminf_uminus[of sequentially "X o r"] assms by auto
+  then show ?thesis by auto
+qed
+
+lemma bounded_abs:
+  assumes "(a::real)<=x" "x<=b"
+  shows "abs x <= max (abs a) (abs b)"
+by (metis abs_less_iff assms leI le_max_iff_disj less_eq_real_def less_le_not_le less_minus_iff minus_minus)
+
+lemma bounded_increasing_convergent2: fixes f::"nat => real"
+  assumes "ALL n. f n <= B"  "ALL n m. n>=m --> f n >= f m"
+  shows "EX l. (f ---> l) sequentially"
+proof-
+def N == "max (abs (f 0)) (abs B)"
+{ fix n have "abs (f n) <= N" unfolding N_def apply (subst bounded_abs) using assms by auto }
+hence "bounded {f n| n::nat. True}" unfolding bounded_real by auto
+from this show ?thesis apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
+   using assms by auto
+qed
+lemma lim_extreal_increasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
+  obtains l where "f ----> (l::extreal)"
+proof(cases "f = (\<lambda>x. - \<infinity>)")
+  case True then show thesis using Lim_const[of "- \<infinity>" sequentially] by (intro that[of "-\<infinity>"]) auto
+next
+  case False
+  from this obtain N where N_def: "f N > (-\<infinity>)" by (auto simp: fun_eq_iff)
+  have "ALL n>=N. f n >= f N" using assms by auto
+  hence minf: "ALL n>=N. f n > (-\<infinity>)" using N_def by auto
+  def Y == "(%n. (if n>=N then f n else f N))"
+  hence incy: "!!n m. n>=m ==> Y n >= Y m" using assms by auto
+  from minf have minfy: "ALL n. Y n ~= (-\<infinity>)" using Y_def by auto
+  show thesis
+  proof(cases "EX B. ALL n. f n < extreal B")
+    case False thus thesis apply- apply(rule that[of \<infinity>]) unfolding Lim_PInfty not_ex not_all
+    apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
+    apply(rule order_trans[OF _ assms[rule_format]]) by auto
+  next case True then guess B ..
+    hence "ALL n. Y n < extreal B" using Y_def by auto note B = this[rule_format]
+    { fix n have "Y n < \<infinity>" using B[of n] apply (subst less_le_trans) by auto
+      hence "Y n ~= \<infinity> & Y n ~= (-\<infinity>)" using minfy by auto
+    } hence *: "ALL n. \<bar>Y n\<bar> \<noteq> \<infinity>" by auto
+    { fix n have "real (Y n) < B" proof- case goal1 thus ?case
+        using B[of n] apply-apply(subst(asm) extreal_real'[THEN sym]) defer defer
+        unfolding extreal_less using * by auto
+      qed
+    }
+    hence B': "ALL n. (real (Y n) <= B)" using less_imp_le by auto
+    have "EX l. (%n. real (Y n)) ----> l"
+      apply(rule bounded_increasing_convergent2)
+    proof safe show "!!n. real (Y n) <= B" using B' by auto
+      fix n m::nat assume "n<=m"
+      hence "extreal (real (Y n)) <= extreal (real (Y m))"
+        using incy[rule_format,of n m] apply(subst extreal_real)+
+        using *[rule_format, of n] *[rule_format, of m] by auto
+      thus "real (Y n) <= real (Y m)" by auto
+    qed then guess l .. note l=this
+    have "Y ----> extreal l" using l apply-apply(subst(asm) lim_extreal[THEN sym])
+    unfolding extreal_real using * by auto
+    thus thesis apply-apply(rule that[of "extreal l"])
+       apply (subst tail_same_limit[of Y _ N]) using Y_def by auto
+  qed
+qed
+
+lemma lim_extreal_decreasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
+  obtains l where "f ----> (l::extreal)"
+proof -
+  from lim_extreal_increasing[of "\<lambda>x. - f x"] assms
+  obtain l where "(\<lambda>x. - f x) ----> l" by auto
+  from extreal_lim_mult[OF this, of "- 1"] show thesis
+    by (intro that[of "-l"]) (simp add: extreal_uminus_eq_reorder)
+qed
+
+lemma compact_extreal:
+  fixes X :: "nat \<Rightarrow> extreal"
+  shows "\<exists>l r. subseq r \<and> (X \<circ> r) ----> l"
+proof -
+  obtain r where "subseq r" and mono: "monoseq (X \<circ> r)"
+    using seq_monosub[of X] unfolding comp_def by auto
+  then have "(\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) m \<le> (X \<circ> r) n) \<or> (\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) n \<le> (X \<circ> r) m)"
+    by (auto simp add: monoseq_def)
+  then obtain l where "(X\<circ>r) ----> l"
+     using lim_extreal_increasing[of "X \<circ> r"] lim_extreal_decreasing[of "X \<circ> r"] by auto
+  then show ?thesis using `subseq r` by auto
+qed
+
+lemma extreal_Sup_lim:
+  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
+  shows "a \<le> Sup s"
+by (metis Lim_bounded_extreal assms complete_lattice_class.Sup_upper)
+
+lemma extreal_Inf_lim:
+  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
+  shows "Inf s \<le> a"
+by (metis Lim_bounded2_extreal assms complete_lattice_class.Inf_lower)
+
+lemma SUP_Lim_extreal:
+  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" "X ----> l" shows "(SUP n. X n) = l"
+proof (rule extreal_SUPI)
+  fix n from assms show "X n \<le> l"
+    by (intro incseq_le_extreal) (simp add: incseq_def)
+next
+  fix y assume "\<And>n. n \<in> UNIV \<Longrightarrow> X n \<le> y"
+  with extreal_Sup_lim[OF _ `X ----> l`, of "{..y}"]
+  show "l \<le> y" by auto
+qed
+
+lemma LIMSEQ_extreal_SUPR:
+  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" shows "X ----> (SUP n. X n)"
+proof (rule lim_extreal_increasing)
+  fix n m :: nat assume "m \<le> n" then show "X m \<le> X n"
+    using `incseq X` by (simp add: incseq_def)
+next
+  fix l assume "X ----> l"
+  with SUP_Lim_extreal[of X, OF assms this] show ?thesis by simp
+qed
+
+lemma INF_Lim_extreal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::extreal)"
+  using SUP_Lim_extreal[of "\<lambda>i. - X i" "- l"]
+  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
+
+lemma LIMSEQ_extreal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: extreal)"
+  using LIMSEQ_extreal_SUPR[of "\<lambda>i. - X i"]
+  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
+
+lemma SUP_eq_LIMSEQ:
+  assumes "mono f"
+  shows "(SUP n. extreal (f n)) = extreal x \<longleftrightarrow> f ----> x"
+proof
+  have inc: "incseq (\<lambda>i. extreal (f i))"
+    using `mono f` unfolding mono_def incseq_def by auto
+  { assume "f ----> x"
+   then have "(\<lambda>i. extreal (f i)) ----> extreal x" by auto
+   from SUP_Lim_extreal[OF inc this]
+   show "(SUP n. extreal (f n)) = extreal x" . }
+  { assume "(SUP n. extreal (f n)) = extreal x"
+    with LIMSEQ_extreal_SUPR[OF inc]
+    show "f ----> x" by auto }
+qed
+
+lemma Liminf_within:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
+proof-
+let ?l="(SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
+{ fix T assume T_def: "open T & mono T & ?l:T"
+  have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T"
+  proof-
+  { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
+  moreover
+  { assume "~(T=UNIV)"
+    then obtain B where "T={B<..}" using T_def extreal_open_mono_set[of T] by auto
+    hence "B<?l" using T_def by auto
+    then obtain d where d_def: "0<d & B<(INF y:(S Int ball x d - {x}). f y)"
+      unfolding less_SUP_iff by auto
+    { fix y assume "y:S & 0 < dist y x & dist y x < d"
+      hence "y:(S Int ball x d - {x})" unfolding ball_def by (auto simp add: dist_commute)
+      hence "f y:T" using d_def INF_leI[of y "S Int ball x d - {x}" f] `T={B<..}` by auto
+    } hence ?thesis apply(rule_tac x="d" in exI) using d_def by auto
+  } ultimately show ?thesis by auto
+  qed
+}
+moreover
+{ fix z
+  assume a: "ALL T. open T --> mono T --> z : T -->
+     (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)"
+  { fix B assume "B<z"
+    then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> B < f y)"
+       using a[rule_format, of "{B<..}"] mono_greaterThan by auto
+    { fix y assume "y:(S Int ball x d - {x})"
+      hence "y:S & 0 < dist y x & dist y x < d" unfolding ball_def apply (simp add: dist_commute)
+         by (metis dist_eq_0_iff real_less_def zero_le_dist)
+      hence "B <= f y" using d_def by auto
+    } hence "B <= INFI (S Int ball x d - {x}) f" apply (subst le_INFI) by auto
+    also have "...<=?l" apply (subst le_SUPI) using d_def by auto
+    finally have "B<=?l" by auto
+  } hence "z <= ?l" using extreal_le_extreal[of z "?l"] by auto
+}
+ultimately show ?thesis unfolding extreal_Liminf_Sup_monoset eventually_within
+   apply (subst extreal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"]) by auto
+qed
+
+lemma Limsup_within:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
+proof-
+let ?l="(INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
+{ fix T assume T_def: "open T & mono (uminus ` T) & ?l:T"
+  have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T"
+  proof-
+  { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
+  moreover
+  { assume "~(T=UNIV)" hence "~(uminus ` T = UNIV)"
+       by (metis Int_UNIV_right Int_absorb1 image_mono extreal_minus_minus_image subset_UNIV)
+    hence "uminus ` T = {Inf (uminus ` T)<..}" using T_def extreal_open_mono_set[of "uminus ` T"]
+       extreal_open_uminus[of T] by auto
+    then obtain B where "T={..<B}"
+      unfolding extreal_Inf_uminus_image_eq extreal_uminus_lessThan[symmetric]
+      unfolding inj_image_eq_iff[OF extreal_inj_on_uminus] by simp
+    hence "?l<B" using T_def by auto
+    then obtain d where d_def: "0<d & (SUP y:(S Int ball x d - {x}). f y)<B"
+      unfolding INF_less_iff by auto
+    { fix y assume "y:S & 0 < dist y x & dist y x < d"
+      hence "y:(S Int ball x d - {x})" unfolding ball_def by (auto simp add: dist_commute)
+      hence "f y:T" using d_def le_SUPI[of y "S Int ball x d - {x}" f] `T={..<B}` by auto
+    } hence ?thesis apply(rule_tac x="d" in exI) using d_def by auto
+  } ultimately show ?thesis by auto
+  qed
+}
+moreover
+{ fix z
+  assume a: "ALL T. open T --> mono (uminus ` T) --> z : T -->
+     (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)"
+  { fix B assume "z<B"
+    then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> f y<B)"
+       using a[rule_format, of "{..<B}"] by auto
+    { fix y assume "y:(S Int ball x d - {x})"
+      hence "y:S & 0 < dist y x & dist y x < d" unfolding ball_def apply (simp add: dist_commute)
+         by (metis dist_eq_0_iff real_less_def zero_le_dist)
+      hence "f y <= B" using d_def by auto
+    } hence "SUPR (S Int ball x d - {x}) f <= B" apply (subst SUP_leI) by auto
+    moreover have "?l<=SUPR (S Int ball x d - {x}) f" apply (subst INF_leI) using d_def by auto
+    ultimately have "?l<=B" by auto
+  } hence "?l <= z" using extreal_ge_extreal[of z "?l"] by auto
+}
+ultimately show ?thesis unfolding extreal_Limsup_Inf_monoset eventually_within
+   apply (subst extreal_InfI) by auto
+qed
+
+
+lemma Liminf_within_UNIV:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Liminf (at x) f = Liminf (at x within UNIV) f"
+by (metis within_UNIV)
+
+
+lemma Liminf_at:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
+using Liminf_within[of x UNIV f] Liminf_within_UNIV[of x f] by auto
+
+
+lemma Limsup_within_UNIV:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Limsup (at x) f = Limsup (at x within UNIV) f"
+by (metis within_UNIV)
+
+
+lemma Limsup_at:
+  fixes f :: "'a::metric_space => extreal"
+  shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
+using Limsup_within[of x UNIV f] Limsup_within_UNIV[of x f] by auto
+
+lemma Lim_within_constant:
+  fixes f :: "'a::metric_space => 'b::topological_space"
+  assumes "ALL y:S. f y = C"
+  shows "(f ---> C) (at x within S)"
+unfolding tendsto_def eventually_within
+by (metis assms(1) linorder_le_less_linear n_not_Suc_n real_of_nat_le_zero_cancel_iff)
+
+lemma Liminf_within_constant:
+  fixes f :: "'a::metric_space => extreal"
+  assumes "ALL y:S. f y = C"
+  assumes "~trivial_limit (at x within S)"
+  shows "Liminf (at x within S) f = C"
+by (metis Lim_within_constant assms lim_imp_Liminf)
+
+lemma Limsup_within_constant:
+  fixes f :: "'a::metric_space => extreal"
+  assumes "ALL y:S. f y = C"
+  assumes "~trivial_limit (at x within S)"
+  shows "Limsup (at x within S) f = C"
+by (metis Lim_within_constant assms lim_imp_Limsup)
+
+lemma islimpt_punctured:
+"x islimpt S = x islimpt (S-{x})"
+unfolding islimpt_def by blast
+
+
+lemma islimpt_in_closure:
+"(x islimpt S) = (x:closure(S-{x}))"
+unfolding closure_def using islimpt_punctured by blast
+
+
+lemma not_trivial_limit_within:
+  "~trivial_limit (at x within S) = (x:closure(S-{x}))"
+using islimpt_in_closure by (metis trivial_limit_within)
+
+
+lemma not_trivial_limit_within_ball:
+  "(~trivial_limit (at x within S)) = (ALL e>0. S Int ball x e - {x} ~= {})"
+  (is "?lhs = ?rhs")
+proof-
+{ assume "?lhs"
+  { fix e :: real assume "e>0"
+    then obtain y where "y:(S-{x}) & dist y x < e"
+       using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto
+    hence "y : (S Int ball x e - {x})" unfolding ball_def by (simp add: dist_commute)
+    hence "S Int ball x e - {x} ~= {}" by blast
+  } hence "?rhs" by auto
+}
+moreover
+{ assume "?rhs"
+  { fix e :: real assume "e>0"
+    then obtain y where "y : (S Int ball x e - {x})" using `?rhs` by blast
+    hence "y:(S-{x}) & dist y x < e" unfolding ball_def by (simp add: dist_commute)
+    hence "EX y:(S-{x}). dist y x < e" by auto
+  } hence "?lhs" using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto
+} ultimately show ?thesis by auto
+qed
+
+subsubsection {* Continuity *}
+
+lemma continuous_imp_tendsto:
+  assumes "continuous (at x0) f"
+  assumes "x ----> x0"
+  shows "(f o x) ----> (f x0)"
+proof-
+{ fix S assume "open S & (f x0):S"
+  from this obtain T where T_def: "open T & x0 : T & (ALL x:T. f x : S)"
+     using assms continuous_at_open by metis
+  hence "(EX N. ALL n>=N. x n : T)" using assms tendsto_explicit T_def by auto
+  hence "(EX N. ALL n>=N. f(x n) : S)" using T_def by auto
+} from this show ?thesis using tendsto_explicit[of "f o x" "f x0"] by auto
+qed
+
+
+lemma continuous_at_sequentially2:
+fixes f :: "'a::metric_space => 'b:: topological_space"
+shows "continuous (at x0) f <-> (ALL x. (x ----> x0) --> (f o x) ----> (f x0))"
+proof-
+{ assume "~(continuous (at x0) f)"
+  from this obtain T where T_def:
+     "open T & f x0 : T & (ALL S. (open S & x0 : S) --> (EX x':S. f x' ~: T))"
+     using continuous_at_open[of x0 f] by metis
+  def X == "{x'. f x' ~: T}" hence "x0 islimpt X" unfolding islimpt_def using T_def by auto
+  from this obtain x where x_def: "(ALL n. x n : X) & x ----> x0"
+     using islimpt_sequential[of x0 X] by auto
+  hence "~(f o x) ----> (f x0)" unfolding tendsto_explicit using X_def T_def by auto
+  hence "EX x. x ----> x0 & (~(f o x) ----> (f x0))" using x_def by auto
+}
+from this show ?thesis using continuous_imp_tendsto by auto
+qed
+
+lemma continuous_at_of_extreal:
+  fixes x0 :: extreal
+  assumes "\<bar>x0\<bar> \<noteq> \<infinity>"
+  shows "continuous (at x0) real"
+proof-
+{ fix T assume T_def: "open T & real x0 : T"
+  def S == "extreal ` T"
+  hence "extreal (real x0) : S" using T_def by auto
+  hence "x0 : S" using assms extreal_real by auto
+  moreover have "open S" using open_extreal S_def T_def by auto
+  moreover have "ALL y:S. real y : T" using S_def T_def by auto
+  ultimately have "EX S. x0 : S & open S & (ALL y:S. real y : T)" by auto
+} from this show ?thesis unfolding continuous_at_open by blast
+qed
+
+
+lemma continuous_at_iff_extreal:
+fixes f :: "'a::t2_space => real"
+shows "continuous (at x0) f <-> continuous (at x0) (extreal o f)"
+proof-
+{ assume "continuous (at x0) f" hence "continuous (at x0) (extreal o f)"
+     using continuous_at_extreal continuous_at_compose[of x0 f extreal] by auto
+}
+moreover
+{ assume "continuous (at x0) (extreal o f)"
+  hence "continuous (at x0) (real o (extreal o f))"
+     using continuous_at_of_extreal by (intro continuous_at_compose[of x0 "extreal o f"]) auto
+  moreover have "real o (extreal o f) = f" using real_extreal_id by (simp add: o_assoc)
+  ultimately have "continuous (at x0) f" by auto
+} ultimately show ?thesis by auto
+qed
+
+
+lemma continuous_on_iff_extreal:
+fixes f :: "'a::t2_space => real"
+fixes A assumes "open A"
+shows "continuous_on A f <-> continuous_on A (extreal o f)"
+   using continuous_at_iff_extreal assms by (auto simp add: continuous_on_eq_continuous_at)
+
+
+lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>)}) real"
+   using continuous_at_of_extreal continuous_on_eq_continuous_at open_image_extreal by auto
+
+
+lemma continuous_on_iff_real:
+  fixes f :: "'a::t2_space => extreal"
+  assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
+  shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)"
+proof-
+  have "f ` A <= UNIV-{\<infinity>,(-\<infinity>)}" using assms by force
+  hence *: "continuous_on (f ` A) real"
+     using continuous_on_real by (simp add: continuous_on_subset)
+have **: "continuous_on ((real o f) ` A) extreal"
+   using continuous_on_extreal continuous_on_subset[of "UNIV" "extreal" "(real o f) ` A"] by blast
+{ assume "continuous_on A f" hence "continuous_on A (real o f)"
+  apply (subst continuous_on_compose) using * by auto
+}
+moreover
+{ assume "continuous_on A (real o f)"
+  hence "continuous_on A (extreal o (real o f))"
+     apply (subst continuous_on_compose) using ** by auto
+  hence "continuous_on A f"
+     apply (subst continuous_on_eq[of A "extreal o (real o f)" f])
+     using assms extreal_real by auto
+}
+ultimately show ?thesis by auto
+qed
+
+
+lemma continuous_at_const:
+  fixes f :: "'a::t2_space => extreal"
+  assumes "ALL x. (f x = C)"
+  shows "ALL x. continuous (at x) f"
+unfolding continuous_at_open using assms t1_space by auto
+
+
+lemma closure_contains_Inf:
+  fixes S :: "real set"
+  assumes "S ~= {}" "EX B. ALL x:S. B<=x"
+  shows "Inf S : closure S"
+proof-
+have *: "ALL x:S. Inf S <= x" using Inf_lower_EX[of _ S] assms by metis
+{ fix e assume "e>(0 :: real)"
+  from this obtain x where x_def: "x:S & x < Inf S + e" using Inf_close `S ~= {}` by auto
+  moreover hence "x > Inf S - e" using * by auto
+  ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff)
+  hence "EX x:S. abs (x - Inf S) < e" using x_def by auto
+} from this show ?thesis apply (subst closure_approachable) unfolding dist_norm by auto
+qed
+
+
+lemma closed_contains_Inf:
+  fixes S :: "real set"
+  assumes "S ~= {}" "EX B. ALL x:S. B<=x"
+  assumes "closed S"
+  shows "Inf S : S"
+by (metis closure_contains_Inf closure_closed assms)
+
+
+lemma mono_closed_real:
+  fixes S :: "real set"
+  assumes mono: "ALL y z. y:S & y<=z --> z:S"
+  assumes "closed S"
+  shows "S = {} | S = UNIV | (EX a. S = {a ..})"
+proof-
+{ assume "S ~= {}"
+  { assume ex: "EX B. ALL x:S. B<=x"
+    hence *: "ALL x:S. Inf S <= x" using Inf_lower_EX[of _ S] ex by metis
+    hence "Inf S : S" apply (subst closed_contains_Inf) using ex `S ~= {}` `closed S` by auto
+    hence "ALL x. (Inf S <= x <-> x:S)" using mono[rule_format, of "Inf S"] * by auto
+    hence "S = {Inf S ..}" by auto
+    hence "EX a. S = {a ..}" by auto
+  }
+  moreover
+  { assume "~(EX B. ALL x:S. B<=x)"
+    hence nex: "ALL B. EX x:S. x<B" by (simp add: not_le)
+    { fix y obtain x where "x:S & x < y" using nex by auto
+      hence "y:S" using mono[rule_format, of x y] by auto
+    } hence "S = UNIV" by auto
+  } ultimately have "S = UNIV | (EX a. S = {a ..})" by blast
+} from this show ?thesis by blast
+qed
+
+
+lemma mono_closed_extreal:
+  fixes S :: "real set"
+  assumes mono: "ALL y z. y:S & y<=z --> z:S"
+  assumes "closed S"
+  shows "EX a. S = {x. a <= extreal x}"
+proof-
+{ assume "S = {}" hence ?thesis apply(rule_tac x=PInfty in exI) by auto }
+moreover
+{ assume "S = UNIV" hence ?thesis apply(rule_tac x="-\<infinity>" in exI) by auto }
+moreover
+{ assume "EX a. S = {a ..}"
+  from this obtain a where "S={a ..}" by auto
+  hence ?thesis apply(rule_tac x="extreal a" in exI) by auto
+} ultimately show ?thesis using mono_closed_real[of S] assms by auto
+qed
+
+subsection {* Sums *}
+
+lemma setsum_extreal[simp]:
+  "(\<Sum>x\<in>A. extreal (f x)) = extreal (\<Sum>x\<in>A. f x)"
+proof cases
+  assume "finite A" then show ?thesis by induct auto
+qed simp
+
+lemma setsum_Pinfty: "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))"
+proof safe
+  assume *: "setsum f P = \<infinity>"
+  show "finite P"
+  proof (rule ccontr) assume "infinite P" with * show False by auto qed
+  show "\<exists>i\<in>P. f i = \<infinity>"
+  proof (rule ccontr)
+    assume "\<not> ?thesis" then have "\<And>i. i \<in> P \<Longrightarrow> f i \<noteq> \<infinity>" by auto
+    from `finite P` this have "setsum f P \<noteq> \<infinity>"
+      by induct auto
+    with * show False by auto
+  qed
+next
+  fix i assume "finite P" "i \<in> P" "f i = \<infinity>"
+  thus "setsum f P = \<infinity>"
+  proof induct
+    case (insert x A)
+    show ?case using insert by (cases "x = i") auto
+  qed simp
+qed
+
+lemma setsum_Inf:
+  shows "\<bar>setsum f A\<bar> = \<infinity> \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>))"
+proof
+  assume *: "\<bar>setsum f A\<bar> = \<infinity>"
+  have "finite A" by (rule ccontr) (insert *, auto)
+  moreover have "\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>"
+  proof (rule ccontr)
+    assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
+    from bchoice[OF this] guess r ..
+    with * show False by (auto simp: setsum_extreal)
+  qed
+  ultimately show "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)" by auto
+next
+  assume "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)"
+  then obtain i where "finite A" "i \<in> A" "\<bar>f i\<bar> = \<infinity>" by auto
+  then show "\<bar>setsum f A\<bar> = \<infinity>"
+  proof induct
+    case (insert j A) then show ?case
+      by (cases rule: extreal3_cases[of "f i" "f j" "setsum f A"]) auto
+  qed simp
+qed
+
+lemma setsum_real_of_extreal:
+  assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
+  shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
+proof -
+  have "\<forall>x\<in>S. \<exists>r. f x = extreal r"
+  proof
+    fix x assume "x \<in> S"
+    from assms[OF this] show "\<exists>r. f x = extreal r" by (cases "f x") auto
+  qed
+  from bchoice[OF this] guess r ..
+  then show ?thesis by simp
+qed
+
+lemma setsum_extreal_0:
+  fixes f :: "'a \<Rightarrow> extreal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
+  shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
+proof
+  assume *: "(\<Sum>x\<in>A. f x) = 0"
+  then have "(\<Sum>x\<in>A. f x) \<noteq> \<infinity>" by auto
+  then have "\<forall>i\<in>A. \<bar>f i\<bar> \<noteq> \<infinity>" using assms by (force simp: setsum_Pinfty)
+  then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
+  from bchoice[OF this] * assms show "\<forall>i\<in>A. f i = 0"
+    using setsum_nonneg_eq_0_iff[of A "\<lambda>i. real (f i)"] by auto
+qed (rule setsum_0')
+
+
+lemma setsum_extreal_right_distrib:
+  fixes f :: "'a \<Rightarrow> extreal" assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
+  shows "r * setsum f A = (\<Sum>n\<in>A. r * f n)"
+proof cases
+  assume "finite A" then show ?thesis using assms
+    by induct (auto simp: extreal_right_distrib setsum_nonneg)
+qed simp
+
+lemma sums_extreal_positive:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
+proof -
+  have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)"
+    using extreal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
+  from LIMSEQ_extreal_SUPR[OF this]
+  show ?thesis unfolding sums_def by (simp add: atLeast0LessThan)
+qed
+
+lemma summable_extreal_pos:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "summable f"
+  using sums_extreal_positive[of f, OF assms] unfolding summable_def by auto
+
+lemma suminf_extreal_eq_SUPR:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i"
+  shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)"
+  using sums_extreal_positive[of f, OF assms, THEN sums_unique] by simp
+
+lemma sums_extreal:
+  "(\<lambda>x. extreal (f x)) sums extreal x \<longleftrightarrow> f sums x"
+  unfolding sums_def by simp
+
+lemma suminf_bound:
+  fixes f :: "nat \<Rightarrow> extreal"
+  assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x" and pos: "\<And>n. 0 \<le> f n"
+  shows "suminf f \<le> x"
+proof (rule Lim_bounded_extreal)
+  have "summable f" using pos[THEN summable_extreal_pos] .
+  then show "(\<lambda>N. \<Sum>n<N. f n) ----> suminf f"
+    by (auto dest!: summable_sums simp: sums_def atLeast0LessThan)
+  show "\<forall>n\<ge>0. setsum f {..<n} \<le> x"
+    using assms by auto
+qed
+
+lemma suminf_bound_add:
+  fixes f :: "nat \<Rightarrow> extreal"
+  assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x" and pos: "\<And>n. 0 \<le> f n" and "y \<noteq> -\<infinity>"
+  shows "suminf f + y \<le> x"
+proof (cases y)
+  case (real r) then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y"
+    using assms by (simp add: extreal_le_minus)
+  then have "(\<Sum> n. f n) \<le> x - y" using pos by (rule suminf_bound)
+  then show "(\<Sum> n. f n) + y \<le> x"
+    using assms real by (simp add: extreal_le_minus)
+qed (insert assms, auto)
+
+lemma sums_finite:
+  assumes "\<forall>N\<ge>n. f N = 0"
+  shows "f sums (\<Sum>N<n. f N)"
+proof -
+  { fix i have "(\<Sum>N<i + n. f N) = (\<Sum>N<n. f N)"
+      by (induct i) (insert assms, auto) }
+  note this[simp]
+  show ?thesis unfolding sums_def
+    by (rule LIMSEQ_offset[of _ n]) (auto simp add: atLeast0LessThan)
+qed
+
+lemma suminf_finite:
+  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add,t2_space}" assumes "\<forall>N\<ge>n. f N = 0"
+  shows "suminf f = (\<Sum>N<n. f N)"
+  using sums_finite[OF assms, THEN sums_unique] by simp
+
+lemma suminf_extreal_0[simp]: "(\<Sum>i. 0) = (0::'a::{comm_monoid_add,t2_space})"
+  using suminf_finite[of 0 "\<lambda>x. 0"] by simp
+
+lemma suminf_upper:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
+  shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)"
+  unfolding suminf_extreal_eq_SUPR[OF assms] SUPR_def
+  by (auto intro: complete_lattice_class.Sup_upper image_eqI)
+
+lemma suminf_0_le:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
+  shows "0 \<le> (\<Sum>n. f n)"
+  using suminf_upper[of f 0, OF assms] by simp
+
+lemma suminf_le_pos:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes "\<And>N. f N \<le> g N" "\<And>N. 0 \<le> f N"
+  shows "suminf f \<le> suminf g"
+proof (safe intro!: suminf_bound)
+  fix n { fix N have "0 \<le> g N" using assms(2,1)[of N] by auto }
+  have "setsum f {..<n} \<le> setsum g {..<n}" using assms by (auto intro: setsum_mono)
+  also have "... \<le> suminf g" using `\<And>N. 0 \<le> g N` by (rule suminf_upper)
+  finally show "setsum f {..<n} \<le> suminf g" .
+qed (rule assms(2))
+
+lemma suminf_half_series_extreal: "(\<Sum>n. (1/2 :: extreal)^Suc n) = 1"
+  using sums_extreal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]
+  by (simp add: one_extreal_def)
+
+lemma suminf_add_extreal:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
+  shows "(\<Sum>i. f i + g i) = suminf f + suminf g"
+  apply (subst (1 2 3) suminf_extreal_eq_SUPR)
+  unfolding setsum_addf
+  by (intro assms extreal_add_nonneg_nonneg SUPR_extreal_add_pos incseq_setsumI setsum_nonneg ballI)+
+
+lemma suminf_cmult_extreal:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes "\<And>i. 0 \<le> f i" "0 \<le> a"
+  shows "(\<Sum>i. a * f i) = a * suminf f"
+  by (auto simp: setsum_extreal_right_distrib[symmetric] assms
+                 extreal_zero_le_0_iff setsum_nonneg suminf_extreal_eq_SUPR
+           intro!: SUPR_extreal_cmult )
+
+lemma suminf_PInfty:
+  assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
+  shows "f i \<noteq> \<infinity>"
+proof -
+  from suminf_upper[of f "Suc i", OF assms(1)] assms(2)
+  have "(\<Sum>i<Suc i. f i) \<noteq> \<infinity>" by auto
+  then show ?thesis
+    unfolding setsum_Pinfty by simp
+qed
+
+lemma suminf_PInfty_fun:
+  assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
+  shows "\<exists>f'. f = (\<lambda>x. extreal (f' x))"
+proof -
+  have "\<forall>i. \<exists>r. f i = extreal r"
+  proof
+    fix i show "\<exists>r. f i = extreal r"
+      using suminf_PInfty[OF assms] assms(1)[of i] by (cases "f i") auto
+  qed
+  from choice[OF this] show ?thesis by auto
+qed
+
+lemma summable_extreal:
+  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
+  shows "summable f"
+proof -
+  have "0 \<le> (\<Sum>i. extreal (f i))"
+    using assms by (intro suminf_0_le) auto
+  with assms obtain r where r: "(\<Sum>i. extreal (f i)) = extreal r"
+    by (cases "\<Sum>i. extreal (f i)") auto
+  from summable_extreal_pos[of "\<lambda>x. extreal (f x)"]
+  have "summable (\<lambda>x. extreal (f x))" using assms by auto
+  from summable_sums[OF this]
+  have "(\<lambda>x. extreal (f x)) sums (\<Sum>x. extreal (f x))" by auto
+  then show "summable f"
+    unfolding r sums_extreal summable_def ..
+qed
+
+lemma suminf_extreal:
+  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
+  shows "(\<Sum>i. extreal (f i)) = extreal (suminf f)"
+proof (rule sums_unique[symmetric])
+  from summable_extreal[OF assms]
+  show "(\<lambda>x. extreal (f x)) sums (extreal (suminf f))"
+    unfolding sums_extreal using assms by (intro summable_sums summable_extreal)
+qed
+
+lemma suminf_extreal_minus:
+  fixes f g :: "nat \<Rightarrow> extreal"
+  assumes ord: "\<And>i. g i \<le> f i" "\<And>i. 0 \<le> g i" and fin: "suminf f \<noteq> \<infinity>" "suminf g \<noteq> \<infinity>"
+  shows "(\<Sum>i. f i - g i) = suminf f - suminf g"
+proof -
+  { fix i have "0 \<le> f i" using ord[of i] by auto }
+  moreover
+  from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] guess f' .. note this[simp]
+  from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] guess g' .. note this[simp]
+  { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: extreal_le_minus_iff) }
+  moreover
+  have "suminf (\<lambda>i. f i - g i) \<le> suminf f"
+    using assms by (auto intro!: suminf_le_pos simp: field_simps)
+  then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>" using fin by auto
+  ultimately show ?thesis using assms `\<And>i. 0 \<le> f i`
+    apply simp
+    by (subst (1 2 3) suminf_extreal)
+       (auto intro!: suminf_diff[symmetric] summable_extreal)
+qed
+
+lemma suminf_extreal_PInf[simp]:
+  "(\<Sum>x. \<infinity>) = \<infinity>"
+proof -
+  have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>)" by (rule suminf_upper) auto
+  then show ?thesis by simp
+qed
+
+lemma summable_real_of_extreal:
+  assumes f: "\<And>i. 0 \<le> f i" and fin: "(\<Sum>i. f i) \<noteq> \<infinity>"
+  shows "summable (\<lambda>i. real (f i))"
+proof (rule summable_def[THEN iffD2])
+  have "0 \<le> (\<Sum>i. f i)" using assms by (auto intro: suminf_0_le)
+  with fin obtain r where r: "extreal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto
+  { fix i have "f i \<noteq> \<infinity>" using f by (intro suminf_PInfty[OF _ fin]) auto
+    then have "\<bar>f i\<bar> \<noteq> \<infinity>" using f[of i] by auto }
+  note fin = this
+  have "(\<lambda>i. extreal (real (f i))) sums (\<Sum>i. extreal (real (f i)))"
+    using f by (auto intro!: summable_extreal_pos summable_sums simp: extreal_le_real_iff zero_extreal_def)
+  also have "\<dots> = extreal r" using fin r by (auto simp: extreal_real)
+  finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_extreal)
+qed
+
+end
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,4 +1,3 @@
-
 header {* Kurzweil-Henstock Gauge Integration in many dimensions. *}
 (*  Author:                     John Harrison
     Translation from HOL light: Robert Himmelmann, TU Muenchen *)
@@ -3780,7 +3779,7 @@
   shows "f x = y"
 proof- have "{x \<in> s. f x \<in> {y}} = {} \<or> {x \<in> s. f x \<in> {y}} = s"
     apply(rule assms(1)[unfolded connected_clopen,rule_format]) apply rule defer
-    apply(rule continuous_closed_in_preimage[OF assms(4) closed_sing])
+    apply(rule continuous_closed_in_preimage[OF assms(4) closed_singleton])
     apply(rule open_openin_trans[OF assms(2)]) unfolding open_contains_ball
   proof safe fix x assume "x\<in>s" 
     from assms(2)[unfolded open_contains_ball,rule_format,OF this] guess e .. note e=conjunctD2[OF this]
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,5 +1,5 @@
 theory Multivariate_Analysis
-imports Fashoda
+imports Fashoda Extended_Real_Limits
 begin
 
 end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -424,80 +424,6 @@
 lemma connected_empty[simp, intro]: "connected {}"
   by (simp add: connected_def)
 
-subsection{* Hausdorff and other separation properties *}
-
-class t0_space = topological_space +
-  assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
-
-class t1_space = topological_space +
-  assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> x \<in> U \<and> y \<notin> U"
-
-instance t1_space \<subseteq> t0_space
-proof qed (fast dest: t1_space)
-
-lemma separation_t1:
-  fixes x y :: "'a::t1_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> x \<in> U \<and> y \<notin> U)"
-  using t1_space[of x y] by blast
-
-lemma closed_sing:
-  fixes a :: "'a::t1_space"
-  shows "closed {a}"
-proof -
-  let ?T = "\<Union>{S. open S \<and> a \<notin> S}"
-  have "open ?T" by (simp add: open_Union)
-  also have "?T = - {a}"
-    by (simp add: set_eq_iff separation_t1, auto)
-  finally show "closed {a}" unfolding closed_def .
-qed
-
-lemma closed_insert [simp]:
-  fixes a :: "'a::t1_space"
-  assumes "closed S" shows "closed (insert a S)"
-proof -
-  from closed_sing assms
-  have "closed ({a} \<union> S)" by (rule closed_Un)
-  thus "closed (insert a S)" by simp
-qed
-
-lemma finite_imp_closed:
-  fixes S :: "'a::t1_space set"
-  shows "finite S \<Longrightarrow> closed S"
-by (induct set: finite, simp_all)
-
-text {* T2 spaces are also known as Hausdorff spaces. *}
-
-class t2_space = topological_space +
-  assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-
-instance t2_space \<subseteq> t1_space
-proof qed (fast dest: hausdorff)
-
-instance metric_space \<subseteq> t2_space
-proof
-  fix x y :: "'a::metric_space"
-  assume xy: "x \<noteq> y"
-  let ?U = "ball x (dist x y / 2)"
-  let ?V = "ball y (dist x y / 2)"
-  have th0: "\<And>d x y z. (d x z :: real) <= d x y + d y z \<Longrightarrow> d y z = d z y
-               ==> ~(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
-  have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
-    using dist_pos_lt[OF xy] th0[of dist,OF dist_triangle dist_commute]
-    by (auto simp add: set_eq_iff)
-  then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-    by blast
-qed
-
-lemma separation_t2:
-  fixes x y :: "'a::t2_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
-  using hausdorff[of x y] by blast
-
-lemma separation_t0:
-  fixes x y :: "'a::t0_space"
-  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
-  using t0_space[of x y] by blast
-
 subsection{* Limit points *}
 
 definition
@@ -994,10 +920,6 @@
 
 text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
 
-definition
-  trivial_limit :: "'a net \<Rightarrow> bool" where
-  "trivial_limit net \<longleftrightarrow> eventually (\<lambda>x. False) net"
-
 lemma trivial_limit_within:
   shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
 proof
@@ -1040,9 +962,6 @@
   apply (simp add: norm_sgn)
   done
 
-lemma trivial_limit_sequentially[intro]: "\<not> trivial_limit sequentially"
-  by (auto simp add: trivial_limit_def eventually_sequentially)
-
 text {* Some property holds "sufficiently close" to the limit point. *}
 
 lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
@@ -1074,6 +993,7 @@
 lemma eventually_False: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
   unfolding trivial_limit_def ..
 
+
 lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
   apply (safe elim!: trivial_limit_eventually)
   apply (simp add: eventually_False [symmetric])
@@ -1417,35 +1337,10 @@
 
 text{* Uniqueness of the limit, when nontrivial. *}
 
-lemma Lim_unique:
-  fixes f :: "'a \<Rightarrow> 'b::t2_space"
-  assumes "\<not> trivial_limit net"  "(f ---> l) net"  "(f ---> l') net"
-  shows "l = l'"
-proof (rule ccontr)
-  assume "l \<noteq> l'"
-  obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
-    using hausdorff [OF `l \<noteq> l'`] by fast
-  have "eventually (\<lambda>x. f x \<in> U) net"
-    using `(f ---> l) net` `open U` `l \<in> U` by (rule topological_tendstoD)
-  moreover
-  have "eventually (\<lambda>x. f x \<in> V) net"
-    using `(f ---> l') net` `open V` `l' \<in> V` by (rule topological_tendstoD)
-  ultimately
-  have "eventually (\<lambda>x. False) net"
-  proof (rule eventually_elim2)
-    fix x
-    assume "f x \<in> U" "f x \<in> V"
-    hence "f x \<in> U \<inter> V" by simp
-    with `U \<inter> V = {}` show "False" by simp
-  qed
-  with `\<not> trivial_limit net` show "False"
-    by (simp add: eventually_False)
-qed
-
 lemma tendsto_Lim:
   fixes f :: "'a \<Rightarrow> 'b::t2_space"
   shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
-  unfolding Lim_def using Lim_unique[of net f] by auto
+  unfolding Lim_def using tendsto_unique[of net f] by auto
 
 text{* Limit under bilinear function *}
 
@@ -1518,7 +1413,7 @@
 apply (rule some_equality)
 apply (rule Lim_at_within)
 apply (rule Lim_ident_at)
-apply (erule Lim_unique [OF assms])
+apply (erule tendsto_unique [OF assms])
 apply (rule Lim_at_within)
 apply (rule Lim_ident_at)
 done
@@ -2558,7 +2453,7 @@
       unfolding islimpt_sequential by auto
     then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
       using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto
-    hence "x \<in> s"  using Lim_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
+    hence "x \<in> s"  using tendsto_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
   }
   thus "closed s" unfolding closed_limpt by auto
 qed
@@ -3131,7 +3026,7 @@
     using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
   moreover
   { fix x assume "P x"
-    hence "l x = l' x" using Lim_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
+    hence "l x = l' x" using tendsto_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
       using l and assms(2) unfolding Lim_sequentially by blast  }
   ultimately show ?thesis by auto
 qed
@@ -5954,7 +5849,7 @@
   hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
     apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
     using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
-  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"]
+  hence "g a = a" using tendsto_unique[OF trivial_limit_sequentially limb, of "g a"]
     unfolding `a=b` and o_assoc by auto
   moreover
   { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
--- a/src/HOL/Probability/Borel_Space.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Borel_Space.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -3,13 +3,9 @@
 header {*Borel spaces*}
 
 theory Borel_Space
-  imports Sigma_Algebra Positive_Extended_Real Multivariate_Analysis
+  imports Sigma_Algebra Multivariate_Analysis
 begin
 
-lemma LIMSEQ_max:
-  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
-  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
-
 section "Generic Borel spaces"
 
 definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = open\<rparr>"
@@ -57,7 +53,7 @@
   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"]
+  from closed_singleton[of "f y"]
   have "{f y} \<in> sets borel" by (rule borel_closed)
   with assms show ?thesis
     unfolding in_borel_measurable_borel `x = f y` by auto
@@ -81,7 +77,7 @@
   shows "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets borel"
   proof (rule borel.insert_in_sets)
     show "{x} \<in> sets borel"
-      using closed_sing[of x] by (rule borel_closed)
+      using closed_singleton[of x] by (rule borel_closed)
   qed simp
 
 lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
@@ -112,26 +108,8 @@
   ultimately show "?I \<in> borel_measurable M" by auto
 qed
 
-lemma borel_measurable_translate:
-  assumes "A \<in> sets borel" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel"
-  shows "f -` A \<in> sets borel"
-proof -
-  have "A \<in> sigma_sets UNIV open" using assms
-    by (simp add: borel_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"
-      using borel.top by simp
-    ultimately show ?case
-      by (auto simp: vimage_Diff borel.Diff)
-  qed (auto simp add: vimage_UN)
-qed
-
 lemma (in sigma_algebra) borel_measurable_restricted:
-  fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
+  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
   shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
     (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
     (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
@@ -142,7 +120,7 @@
   show ?thesis unfolding *
     unfolding in_borel_measurable_borel
   proof (simp, safe)
-    fix S :: "'x set" assume "S \<in> sets borel"
+    fix S :: "extreal set" assume "S \<in> sets borel"
       "\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
     then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
     then have f: "?f -` S \<inter> A \<in> sets M"
@@ -161,7 +139,7 @@
       then show ?thesis using f by auto
     qed
   next
-    fix S :: "'x set" assume "S \<in> sets borel"
+    fix S :: "extreal set" assume "S \<in> sets borel"
       "\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M"
     then have f: "?f -` S \<inter> space M \<in> sets M" by auto
     then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
@@ -1024,103 +1002,6 @@
   using borel_measurable_euclidean_component
   unfolding nth_conv_component by auto
 
-section "Borel space over the real line with infinity"
-
-lemma borel_Real_measurable:
-  "A \<in> sets borel \<Longrightarrow> Real -` A \<in> sets borel"
-proof (rule borel_measurable_translate)
-  fix B :: "pextreal 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_pextreal_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"
-    using `open T` by auto
-qed simp
-
-lemma borel_real_measurable:
-  "A \<in> sets borel \<Longrightarrow> (real -` A :: pextreal set) \<in> sets borel"
-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 <..}) :: pextreal set)"
-    unfolding open_pextreal_def using `open B`
-    by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
-  then show "(real -` B :: pextreal set) \<in> sets borel" 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
-proof safe
-  fix S :: "pextreal set" assume "S \<in> sets borel"
-  from borel_Real_measurable[OF this]
-  have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
-    using assms
-    unfolding vimage_compose in_borel_measurable_borel
-    by auto
-  thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
-  unfolding in_borel_measurable_borel
-proof safe
-  fix S :: "real set" assume "S \<in> sets borel"
-  from borel_real_measurable[OF this]
-  have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
-    using assms
-    unfolding vimage_compose in_borel_measurable_borel
-    by auto
-  thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-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 (in sigma_algebra) borel_measurable_pextreal_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: fun_eq_iff Real_real)
-  show "f \<in> borel_measurable M"
-    apply (subst f)
-    apply (rule measurable_If)
-    using * ** by auto
-qed
-
 lemma borel_measurable_continuous_on1:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
   assumes "continuous_on UNIV f"
@@ -1187,206 +1068,213 @@
   using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]]
   by (simp add: comp_def)
 
+subsection "Borel space on the extended reals"
+
+lemma borel_measurable_extreal_borel:
+  "extreal \<in> borel_measurable borel"
+  unfolding borel_def[where 'a=extreal]
+proof (rule borel.measurable_sigma)
+  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
+  then have "open X" by (auto simp: mem_def)
+  then have "open (extreal -` X \<inter> space borel)"
+    by (simp add: open_extreal_vimage)
+  then show "extreal -` X \<inter> space borel \<in> sets borel" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_extreal[simp, intro]:
+  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
+  using measurable_comp[OF f borel_measurable_extreal_borel] unfolding comp_def .
+
+lemma borel_measurable_real_of_extreal_borel:
+  "(real :: extreal \<Rightarrow> real) \<in> borel_measurable borel"
+  unfolding borel_def[where 'a=real]
+proof (rule borel.measurable_sigma)
+  fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
+  then have "open B" by (auto simp: mem_def)
+  have *: "extreal -` real -` (B - {0}) = B - {0}" by auto
+  have open_real: "open (real -` (B - {0}) :: extreal set)"
+    unfolding open_extreal_def * using `open B` by auto
+  show "(real -` B \<inter> space borel :: extreal set) \<in> sets borel"
+  proof cases
+    assume "0 \<in> B"
+    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0}"
+      by (auto simp add: real_of_extreal_eq_0)
+    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+      using open_real by auto
+  next
+    assume "0 \<notin> B"
+    then have *: "(real -` B :: extreal set) = real -` (B - {0})"
+      by (auto simp add: real_of_extreal_eq_0)
+    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+      using open_real by auto
+  qed
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_real_of_extreal[simp, intro]:
+  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: extreal)) \<in> borel_measurable M"
+  using measurable_comp[OF f borel_measurable_real_of_extreal_borel] unfolding comp_def .
+
+lemma (in sigma_algebra) borel_measurable_extreal_iff:
+  shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+  assume "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
+  from borel_measurable_real_of_extreal[OF this]
+  show "f \<in> borel_measurable M" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_iff_real:
+  "f \<in> borel_measurable M \<longleftrightarrow>
+    ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
+proof safe
+  assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
+  have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
+  with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
+  let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else extreal (real (f x))"
+  have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
+  also have "?f = f" by (auto simp: fun_eq_iff extreal_real)
+  finally show "f \<in> borel_measurable M" .
+qed (auto intro: measurable_sets borel_measurable_real_of_extreal)
 
 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"
+  shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<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
+  assume "f -` {a <..} \<inter> space M \<in> sets M"
+  moreover have "f -` {..a} \<inter> space M = space M - f -` {a <..} \<inter> space M" by auto
+  ultimately show "f -` {..a} \<inter> space M \<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
+  assume "f -` {..a} \<inter> space M \<in> sets M"
+  moreover have "f -` {a <..} \<inter> space M = space M - f -` {..a} \<inter> space M" by auto
+  ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
 qed
 
 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"
+  shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<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
+  assume "f -` {a ..} \<inter> space M \<in> sets M"
+  moreover have "f -` {..< a} \<inter> space M = space M - f -` {a ..} \<inter> space M" by auto
+  ultimately show "f -` {..< a} \<inter> space M \<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
+  assume "f -` {..< a} \<inter> space M \<in> sets M"
+  moreover have "f -` {a ..} \<inter> space M = space M - f -` {..< a} \<inter> space M" by auto
+  ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
 qed
 
-lemma (in sigma_algebra) less_eq_le_pextreal_measurable:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  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)"
+lemma (in sigma_algebra) borel_measurable_uminus_borel_extreal:
+  "(uminus :: extreal \<Rightarrow> extreal) \<in> borel_measurable borel"
+proof (subst borel_def, rule borel.measurable_sigma)
+  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
+  then have "open X" by (simp add: mem_def)
+  have "uminus -` X = uminus ` X" by (force simp: image_iff)
+  then have "open (uminus -` X)" using `open X` extreal_open_uminus by auto
+  then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_uminus_extreal[intro]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M"
+  using measurable_comp[OF assms borel_measurable_uminus_borel_extreal] by (simp add: comp_def)
+
+lemma (in sigma_algebra) borel_measurable_uminus_eq_extreal[simp]:
+  "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
 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 x assume "a < f x" and [simp]: "x \<in> space M"
-        with ex_pextreal_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: pextreal_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 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
-      with a show ?thesis by auto
-    qed simp
+  assume ?l from borel_measurable_uminus_extreal[OF this] show ?r by simp
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_eq_atMost_extreal:
+  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
+proof (intro iffI allI)
+  assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
+  show "f \<in> borel_measurable M"
+    unfolding borel_measurable_extreal_iff_real borel_measurable_iff_le
+  proof (intro conjI allI)
+    fix a :: real
+    { fix x :: extreal assume *: "\<forall>i::nat. real i < x"
+      have "x = \<infinity>"
+      proof (rule extreal_top)
+        fix B from real_arch_lt[of B] guess n ..
+        then have "extreal B < real n" by auto
+        with * show "B \<le> x" by (metis less_trans less_imp_le)
+      qed }
+    then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
+      by (auto simp: not_le)
+    then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
+    moreover
+    have "{-\<infinity>} = {..-\<infinity>}" by auto
+    then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
+    moreover have "{x\<in>space M. f x \<le> extreal a} \<in> sets M"
+      using pos[of "extreal a"] by (simp add: vimage_def Int_def conj_commute)
+    moreover have "{w \<in> space M. real (f w) \<le> a} =
+      (if a < 0 then {w \<in> space M. f w \<le> extreal a} - f -` {-\<infinity>} \<inter> space M
+      else {w \<in> space M. f w \<le> extreal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
+      proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed
+    ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
   qed
-next
-  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_pextreal_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: pextreal_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
+qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_greater:
-  "(f::'a \<Rightarrow> pextreal) \<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_pextreal_measurable
-    unfolding greater_eq_le_measurable .
-  show "f \<in> borel_measurable M" unfolding borel_measurable_pextreal_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
-    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 (in sigma_algebra) borel_measurable_eq_atLeast_extreal:
+  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
+proof
+  assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
+  moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
+    by (auto simp: extreal_uminus_le_reorder)
+  ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
+    unfolding borel_measurable_eq_atMost_extreal by auto
+  then show "f \<in> borel_measurable M" by simp
+qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_less:
-  "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
-  using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable greater_eq_le_measurable .
+lemma (in sigma_algebra) borel_measurable_extreal_iff_less:
+  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
+  unfolding borel_measurable_eq_atLeast_extreal greater_eq_le_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_le:
-  "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
-  using borel_measurable_pextreal_iff_greater unfolding less_eq_ge_measurable .
+lemma (in sigma_algebra) borel_measurable_extreal_iff_ge:
+  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
+  unfolding borel_measurable_eq_atMost_extreal less_eq_ge_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_pextreal_iff_ge:
-  "(f::'a \<Rightarrow> pextreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
-  using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable .
-
-lemma (in sigma_algebra) borel_measurable_pextreal_eq_const:
-  fixes f :: "'a \<Rightarrow> pextreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_extreal_eq_const:
+  fixes f :: "'a \<Rightarrow> extreal" 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_pextreal_neq_const:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_extreal_neq_const:
+  fixes f :: "'a \<Rightarrow> extreal" 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_pextreal_less[intro,simp]:
-  fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_le[intro,simp]:
+  fixes f g :: "'a \<Rightarrow> extreal"
+  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} =
+    {x \<in> space M. real (f x) \<le> real (g x)} - (f -` {\<infinity>, -\<infinity>} \<inter> space M \<union> g -` {\<infinity>, -\<infinity>} \<inter> space M) \<union>
+    f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
+  proof (intro set_eqI)
+    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: extreal2_cases[of "f x" "g x"]) auto
+  qed
+  with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_extreal_less[intro,simp]:
+  fixes f :: "'a \<Rightarrow> extreal"
   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_pextreal_neq_const)
-  moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pextreal_eq_const)
-  moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pextreal_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_pextreal_strict_mono_iff)
-  ultimately show ?thesis by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_pextreal_le[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  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
+  have "{x \<in> space M. f x < g x} = space M - {x \<in> space M. g x \<le> f x}" by auto
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_eq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_eq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> extreal"
   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"
@@ -1395,8 +1283,8 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_neq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_neq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> extreal"
   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"
@@ -1405,20 +1293,28 @@
   thus ?thesis using f g by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_add[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) split_sets:
+  "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
+  "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
+  by auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_add[intro, simp]:
+  fixes f :: "'a \<Rightarrow> extreal"
   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 = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
-     by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex)
-  show ?thesis using assms unfolding *
-    by (auto intro!: measurable_If)
+  { fix x assume "x \<in> space M" then have "f x + g x =
+      (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
+        else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>
+        else extreal (real (f x) + real (g x)))"
+      by (cases rule: extreal2_cases[of "f x" "g x"]) auto }
+  with assms show ?thesis
+    by (auto cong: measurable_cong simp: split_sets
+             intro!: Un measurable_If measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_setsum[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
   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
@@ -1427,20 +1323,49 @@
     by induct auto
 qed (simp add: borel_measurable_const)
 
-lemma (in sigma_algebra) borel_measurable_pextreal_times[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pextreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+lemma abs_extreal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: extreal\<bar> = x"
+  by (cases x) auto
+
+lemma abs_extreal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: extreal\<bar> = -x"
+  by (cases x) auto
+
+lemma abs_extreal_pos[simp]: "0 \<le> \<bar>x :: extreal\<bar>"
+  by (cases x) auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_abs[intro, simp]:
+  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
+proof -
+  { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
+  then show ?thesis using assms by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_extreal_times[intro, simp]:
+  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
 proof -
+  { fix f g :: "'a \<Rightarrow> extreal"
+    assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+      and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x"
+    { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
+        else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
+        else extreal (real (f x) * real (g x)))"
+      apply (cases rule: extreal2_cases[of "f x" "g x"])
+      using pos[of x] by auto }
+    with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+      by (auto cong: measurable_cong simp: split_sets
+               intro!: Un measurable_If measurable_sets) }
+  note pos_times = this
   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: fun_eq_iff pextreal_noteq_omega_Ex)
+    (\<lambda>x. if 0 \<le> f x \<and> 0 \<le> g x \<or> f x < 0 \<and> g x < 0 then \<bar>f x\<bar> * \<bar>g x\<bar> else - (\<bar>f x\<bar> * \<bar>g x\<bar>))"
+    by (auto simp: fun_eq_iff)
   show ?thesis using assms unfolding *
-    by (auto intro!: measurable_If)
+    by (intro measurable_If pos_times borel_measurable_uminus_extreal)
+       (auto simp: split_sets intro!: Int)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_setprod[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_setprod[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
 proof cases
@@ -1448,64 +1373,73 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma (in sigma_algebra) borel_measurable_pextreal_min[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_min[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> extreal"
   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_pextreal_max[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_extreal_max[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> extreal"
   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> pextreal"
+  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> extreal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
-  unfolding borel_measurable_pextreal_iff_greater
-proof safe
+  unfolding borel_measurable_extreal_iff_ge
+proof
   fix a
-  have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
+  have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
     by (auto simp: less_SUP_iff SUPR_apply)
-  then show "{x\<in>space M. a < ?sup x} \<in> sets M"
+  then show "?sup -` {a<..} \<inter> space M \<in> sets M"
     using assms by auto
 qed
 
 lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
-  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pextreal"
+  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> extreal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
-  unfolding borel_measurable_pextreal_iff_less
-proof safe
+  unfolding borel_measurable_extreal_iff_less
+proof
   fix a
-  have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
+  have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
     by (auto simp: INF_less_iff INFI_apply)
-  then show "{x\<in>space M. ?inf x < a} \<in> sets M"
+  then show "?inf -` {..<a} \<inter> space M \<in> sets M"
     using assms by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pextreal_diff[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pextreal"
+lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  assumes "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+  unfolding liminf_SUPR_INFI using assms by auto
+
+lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  assumes "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
+  unfolding limsup_INFI_SUPR using assms by auto
+
+lemma (in sigma_algebra) borel_measurable_extreal_diff[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> extreal"
   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_pextreal_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: pextreal_less_minus_iff)
-  then show "{x \<in> space M. a < f x - g x} \<in> sets M"
-    using assms by auto
-qed
+  unfolding minus_extreal_def using assms by auto
 
 lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
-  using assms unfolding psuminf_def by auto
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
+  shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
+  apply (subst measurable_cong)
+  apply (subst suminf_extreal_eq_SUPR)
+  apply (rule pos)
+  using assms by auto
 
 section "LIMSEQ is borel measurable"
 
@@ -1515,28 +1449,11 @@
   and u: "\<And>i. u i \<in> borel_measurable M"
   shows "u' \<in> borel_measurable M"
 proof -
-  let "?pu x i" = "max (u i x) 0"
-  let "?nu x i" = "max (- u i x) 0"
-  { fix x assume x: "x \<in> space M"
-    have "(?pu x) ----> max (u' x) 0"
-      "(?nu x) ----> max (- u' x) 0"
-      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
-    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
-    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
-      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
-      by (simp_all add: Real_max'[symmetric]) }
-  note eq = this
-  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. extreal (u n x)) = extreal (u' x)"
+    using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_extreal)
+  moreover from u have "(\<lambda>x. liminf (\<lambda>n. extreal (u n x))) \<in> borel_measurable M"
     by auto
-  have "(\<lambda>x. SUP n. INF m. Real (u (n + m) x)) \<in> borel_measurable M"
-       "(\<lambda>x. SUP n. INF m. Real (- u (n + m) x)) \<in> borel_measurable M"
-    using u by auto
-  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel]
-  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
-       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M" by auto
-  note this[THEN borel_measurable_real]
-  from borel_measurable_diff[OF this]
-  show ?thesis unfolding * .
+  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_extreal_iff)
 qed
 
 end
--- a/src/HOL/Probability/Caratheodory.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Caratheodory.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,36 +1,66 @@
 header {*Caratheodory Extension Theorem*}
 
 theory Caratheodory
-  imports Sigma_Algebra Positive_Extended_Real
+  imports Sigma_Algebra Extended_Real_Limits
 begin
 
+lemma suminf_extreal_2dimen:
+  fixes f:: "nat \<times> nat \<Rightarrow> extreal"
+  assumes pos: "\<And>p. 0 \<le> f p"
+  assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
+  shows "(\<Sum>i. f (prod_decode i)) = suminf g"
+proof -
+  have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
+    using assms by (simp add: fun_eq_iff)
+  have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
+    by (simp add: setsum_reindex[OF inj_prod_decode] comp_def)
+  { fix n
+    let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
+    { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
+      then have "a < ?M fst" "b < ?M snd"
+        by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
+    then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
+      by (auto intro!: setsum_mono3 simp: pos)
+    then have "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto }
+  moreover
+  { fix a b
+    let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
+    { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
+        by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
+    then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
+      by (auto intro!: setsum_mono3 simp: pos) }
+  ultimately
+  show ?thesis unfolding g_def using pos
+    by (auto intro!: SUPR_eq  simp: setsum_cartesian_product reindex le_SUPI2
+                     setsum_nonneg suminf_extreal_eq_SUPR SUPR_pair
+                     SUPR_extreal_setsum[symmetric] incseq_setsumI setsum_nonneg)
+qed
+
 text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
 
 subsection {* Measure Spaces *}
 
 record 'a measure_space = "'a algebra" +
-  measure :: "'a set \<Rightarrow> pextreal"
+  measure :: "'a set \<Rightarrow> extreal"
 
-definition positive where "positive M f \<longleftrightarrow> f {} = (0::pextreal)"
-  -- "Positive is enforced by the type"
+definition positive where "positive M f \<longleftrightarrow> f {} = (0::extreal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
 
 definition additive where "additive M f \<longleftrightarrow>
   (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
 
-definition countably_additive where "countably_additive M f \<longleftrightarrow>
-  (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
-    (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i))"
+definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> extreal) \<Rightarrow> bool" where
+  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
+    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
 
 definition increasing where "increasing M f \<longleftrightarrow>
   (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<subseteq> y \<longrightarrow> f x \<le> f y)"
 
 definition subadditive where "subadditive M f \<longleftrightarrow>
-  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow>
-    f (x \<union> y) \<le> f x + f y)"
+  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
 
 definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow>
   (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
-    f (\<Union>i. A i) \<le> (\<Sum>\<^isub>\<infinity> n. f (A n)))"
+    (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
 
 definition lambda_system where "lambda_system M f = {l \<in> sets M.
   \<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x}"
@@ -39,14 +69,19 @@
   positive M f \<and> increasing M f \<and> countably_subadditive M f"
 
 definition measure_set where "measure_set M f X = {r.
-  \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>\<^isub>\<infinity> i. f (A i)) = r}"
+  \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
 
 locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" +
-  assumes empty_measure [simp]: "measure M {} = 0"
+  assumes measure_positive: "positive M (measure M)"
       and ca: "countably_additive M (measure M)"
 
 abbreviation (in measure_space) "\<mu> \<equiv> measure M"
 
+lemma (in measure_space)
+  shows empty_measure[simp, intro]: "\<mu> {} = 0"
+  and positive_measure[simp, intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> 0 \<le> \<mu> A"
+  using measure_positive unfolding positive_def by auto
+
 lemma increasingD:
   "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
   by (auto simp add: increasing_def)
@@ -61,39 +96,30 @@
     \<Longrightarrow> f (x \<union> y) = f x + f y"
   by (auto simp add: additive_def)
 
-lemma countably_additiveD:
-  "countably_additive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A
-    \<Longrightarrow> (\<Union>i. A i) \<in> sets M \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)"
-  by (simp add: countably_additive_def)
-
-lemma countably_subadditiveD:
-  "countably_subadditive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow>
-   (\<Union>i. A i) \<in> sets M \<Longrightarrow> f (\<Union>i. A i) \<le> psuminf (f o A)"
-  by (auto simp add: countably_subadditive_def o_def)
-
 lemma countably_additiveI:
-  "(\<And>A. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
-    \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)) \<Longrightarrow> countably_additive M f"
-  by (simp add: countably_additive_def)
+  assumes "\<And>A x. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
+    \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
+  shows "countably_additive M f"
+  using assms by (simp add: countably_additive_def)
 
 section "Extend binary sets"
 
 lemma LIMSEQ_binaryset:
   assumes f: "f {} = 0"
-  shows  "(\<lambda>n. \<Sum>i = 0..<n. f (binaryset A B i)) ----> f A + f B"
+  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
 proof -
-  have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
+  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
     proof
       fix n
-      show "(\<Sum>i\<Colon>nat = 0\<Colon>nat..<Suc (Suc n). f (binaryset A B i)) = f A + f B"
+      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
         by (induct n)  (auto simp add: binaryset_def f)
     qed
   moreover
   have "... ----> f A + f B" by (rule LIMSEQ_const)
   ultimately
-  have "(\<lambda>n. \<Sum>i = 0..< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
+  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
     by metis
-  hence "(\<lambda>n. \<Sum>i = 0..< n+2. f (binaryset A B i)) ----> f A + f B"
+  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
     by simp
   thus ?thesis by (rule LIMSEQ_offset [where k=2])
 qed
@@ -101,28 +127,13 @@
 lemma binaryset_sums:
   assumes f: "f {} = 0"
   shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
-    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
+    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
 
 lemma suminf_binaryset_eq:
-  fixes f :: "'a set \<Rightarrow> real"
+  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
   shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
   by (metis binaryset_sums sums_unique)
 
-lemma binaryset_psuminf:
-  assumes "f {} = 0"
-  shows "(\<Sum>\<^isub>\<infinity> n. f (binaryset A B n)) = f A + f B" (is "?suminf = ?sum")
-proof -
-  have *: "{..<2} = {0, 1::nat}" by auto
-  have "\<forall>n\<ge>2. f (binaryset A B n) = 0"
-    unfolding binaryset_def
-    using assms by auto
-  hence "?suminf = (\<Sum>N<2. f (binaryset A B N))"
-    by (rule psuminf_finite)
-  also have "... = ?sum" unfolding * binaryset_def
-    by simp
-  finally show ?thesis .
-qed
-
 subsection {* Lambda Systems *}
 
 lemma (in algebra) lambda_system_eq:
@@ -144,7 +155,7 @@
   by (simp add: lambda_system_def)
 
 lemma (in algebra) lambda_system_Compl:
-  fixes f:: "'a set \<Rightarrow> pextreal"
+  fixes f:: "'a set \<Rightarrow> extreal"
   assumes x: "x \<in> lambda_system M f"
   shows "space M - x \<in> lambda_system M f"
 proof -
@@ -157,7 +168,7 @@
 qed
 
 lemma (in algebra) lambda_system_Int:
-  fixes f:: "'a set \<Rightarrow> pextreal"
+  fixes f:: "'a set \<Rightarrow> extreal"
   assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
   shows "x \<inter> y \<in> lambda_system M f"
 proof -
@@ -191,7 +202,7 @@
 qed
 
 lemma (in algebra) lambda_system_Un:
-  fixes f:: "'a set \<Rightarrow> pextreal"
+  fixes f:: "'a set \<Rightarrow> extreal"
   assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
   shows "x \<union> y \<in> lambda_system M f"
 proof -
@@ -250,53 +261,54 @@
     by (auto simp add: disjoint_family_on_def binaryset_def)
   hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
          (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
-    using cs by (simp add: countably_subadditive_def)
+         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
+    using cs by (auto simp add: countably_subadditive_def)
   hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
-         f (x \<union> y) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
     by (simp add: range_binaryset_eq UN_binaryset_eq)
   thus "f (x \<union> y) \<le>  f x + f y" using f x y
-    by (auto simp add: Un o_def binaryset_psuminf positive_def)
+    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
 qed
 
 lemma (in algebra) additive_sum:
   fixes A:: "nat \<Rightarrow> 'a set"
-  assumes f: "positive M f" and ad: "additive M f"
+  assumes f: "positive M f" and ad: "additive M f" and "finite S"
       and A: "range A \<subseteq> sets M"
-      and disj: "disjoint_family A"
-  shows  "setsum (f \<circ> A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
-proof (induct n)
-  case 0 show ?case using f by (simp add: positive_def)
+      and disj: "disjoint_family_on A S"
+  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
+using `finite S` disj proof induct
+  case empty show ?case using f by (simp add: positive_def)
 next
-  case (Suc n)
-  have "A n \<inter> (\<Union>i\<in>{0..<n}. A i) = {}" using disj
-    by (auto simp add: disjoint_family_on_def neq_iff) blast
+  case (insert s S)
+  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
+    by (auto simp add: disjoint_family_on_def neq_iff)
   moreover
-  have "A n \<in> sets M" using A by blast
-  moreover have "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
-    by (metis A UNION_in_sets atLeast0LessThan)
+  have "A s \<in> sets M" using A by blast
+  moreover have "(\<Union>i\<in>S. A i) \<in> sets M"
+    using A `finite S` by auto
   moreover
-  ultimately have "f (A n \<union> (\<Union>i\<in>{0..<n}. A i)) = f (A n) + f(\<Union>i\<in>{0..<n}. A i)"
+  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
     using ad UNION_in_sets A by (auto simp add: additive_def)
-  with Suc.hyps show ?case using ad
-    by (auto simp add: atLeastLessThanSuc additive_def)
+  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
+    by (auto simp add: additive_def subset_insertI)
 qed
 
 lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> pextreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> extreal"
   assumes f: "positive M f" and ad: "additive M f"
       and inc: "increasing M f"
       and A: "range A \<subseteq> sets M"
       and disj: "disjoint_family A"
-  shows  "psuminf (f \<circ> A) \<le> f (space M)"
-proof (safe intro!: psuminf_bound)
+  shows  "(\<Sum>i. f (A i)) \<le> f (space M)"
+proof (safe intro!: suminf_bound)
   fix N
-  have "setsum (f \<circ> A) {0..<N} = f (\<Union>i\<in>{0..<N}. A i)"
-    by (rule additive_sum [OF f ad A disj])
+  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
+  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
+    by (rule additive_sum [OF f ad _ A]) (auto simp: disj_N)
   also have "... \<le> f (space M)" using space_closed A
-    by (blast intro: increasingD [OF inc] UNION_in_sets top)
-  finally show "setsum (f \<circ> A) {..<N} \<le> f (space M)" by (simp add: atLeast0LessThan)
-qed
+    by (intro increasingD[OF inc] finite_UN) auto
+  finally show "(\<Sum>i<N. f (A i)) \<le> f (space M)" by simp
+qed (insert f A, auto simp: positive_def)
 
 lemma lambda_system_increasing:
  "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
@@ -307,7 +319,7 @@
   by (simp add: positive_def lambda_system_def)
 
 lemma (in algebra) lambda_system_strong_sum:
-  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pextreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> extreal"
   assumes f: "positive M f" and a: "a \<in> sets M"
       and A: "range A \<subseteq> lambda_system M f"
       and disj: "disjoint_family A"
@@ -331,7 +343,7 @@
   assumes oms: "outer_measure_space M f"
       and A: "range A \<subseteq> lambda_system M f"
       and disj: "disjoint_family A"
-  shows  "(\<Union>i. A i) \<in> lambda_system M f \<and> psuminf (f \<circ> A) = f (\<Union>i. A i)"
+  shows  "(\<Union>i. A i) \<in> lambda_system M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
 proof -
   have pos: "positive M f" and inc: "increasing M f"
    and csa: "countably_subadditive M f"
@@ -347,15 +359,16 @@
 
   have U_in: "(\<Union>i. A i) \<in> sets M"
     by (metis A'' countable_UN)
-  have U_eq: "f (\<Union>i. A i) = psuminf (f o A)"
+  have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))"
   proof (rule antisym)
-    show "f (\<Union>i. A i) \<le> psuminf (f \<circ> A)"
-      by (rule countably_subadditiveD [OF csa A'' disj U_in])
-    show "psuminf (f \<circ> A) \<le> f (\<Union>i. A i)"
-      by (rule psuminf_bound, unfold atLeast0LessThan[symmetric])
-         (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right
-                lambda_system_positive lambda_system_additive
-                subset_Un_eq increasingD [OF inc] A' A'' UNION_in_sets U_in)
+    show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))"
+      using csa[unfolded countably_subadditive_def] A'' disj U_in by auto
+    have *: "\<And>i. 0 \<le> f (A i)" using pos A'' unfolding positive_def by auto
+    have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
+    show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
+      using algebra.additive_sum [OF alg_ls lambda_system_positive[OF pos] lambda_system_additive _ A' dis]
+      using A''
+      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN)
   qed
   {
     fix a
@@ -373,15 +386,15 @@
         have "a \<inter> (\<Union>i. A i) \<in> sets M"
           by (metis Int U_in a)
         ultimately
-        have "f (a \<inter> (\<Union>i. A i)) \<le> psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)"
-          using countably_subadditiveD [OF csa, of "(\<lambda>i. a \<inter> A i)"]
+        have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
+          using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"]
           by (simp add: o_def)
         hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le>
-            psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i))"
+            (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
           by (rule add_right_mono)
         moreover
-        have "psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i)) \<le> f a"
-          proof (safe intro!: psuminf_bound_add)
+        have "(\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+          proof (intro suminf_bound_add allI)
             fix n
             have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
               by (metis A'' UNION_in_sets)
@@ -395,8 +408,14 @@
             have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
               by (blast intro: increasingD [OF inc] UNION_eq_Union_image
                                UNION_in U_in)
-            thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {..<n} + f (a - (\<Union>i. A i)) \<le> f a"
+            thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
               by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
+          next
+            have "\<And>i. a \<inter> A i \<in> sets M" using A'' by auto
+            then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
+            have "\<And>i. a - (\<Union>i. A i) \<in> sets M" using A'' by auto
+            then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
+            then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
           qed
         ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
           by (rule order_trans)
@@ -443,12 +462,14 @@
 proof (auto simp add: increasing_def)
   fix x y
   assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
-  have "f x \<le> f x + f (y-x)" ..
+  then have "y - x \<in> sets M" by auto
+  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
+  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
   also have "... = f (x \<union> (y-x))" using addf
     by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
   also have "... = f y"
     by (metis Un_Diff_cancel Un_absorb1 xy(3))
-  finally show "f x \<le> f y" .
+  finally show "f x \<le> f y" by simp
 qed
 
 lemma (in algebra) countably_additive_additive:
@@ -461,27 +482,27 @@
     by (auto simp add: disjoint_family_on_def binaryset_def)
   hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
          (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
     using ca
     by (simp add: countably_additive_def)
   hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
-         f (x \<union> y) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
+         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
     by (simp add: range_binaryset_eq UN_binaryset_eq)
   thus "f (x \<union> y) = f x + f y" using posf x y
-    by (auto simp add: Un binaryset_psuminf positive_def)
+    by (auto simp add: Un suminf_binaryset_eq positive_def)
 qed
 
 lemma inf_measure_nonempty:
   assumes f: "positive M f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
   shows "f b \<in> measure_set M f a"
 proof -
-  have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
-    by (rule psuminf_finite) (simp add: f[unfolded positive_def])
+  let ?A = "\<lambda>i::nat. (if i = 0 then b else {})"
+  have "(\<Sum>i. f (?A i)) = (\<Sum>i<1::nat. f (?A i))"
+    by (rule suminf_finite) (simp add: f[unfolded positive_def])
   also have "... = f b"
     by simp
-  finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
-  thus ?thesis using assms
-    by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
+  finally show ?thesis using assms
+    by (auto intro!: exI [of _ ?A]
              simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
 qed
 
@@ -489,19 +510,19 @@
   assumes posf: "positive M f" and ca: "countably_additive M f"
       and s: "s \<in> sets M"
   shows "Inf (measure_set M f s) = f s"
-  unfolding Inf_pextreal_def
+  unfolding Inf_extreal_def
 proof (safe intro!: Greatest_equality)
   fix z
   assume z: "z \<in> measure_set M f s"
   from this obtain A where
     A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
-    and "s \<subseteq> (\<Union>x. A x)" and si: "psuminf (f \<circ> A) = z"
+    and "s \<subseteq> (\<Union>x. A x)" and si: "(\<Sum>i. f (A i)) = z"
     by (auto simp add: measure_set_def comp_def)
   hence seq: "s = (\<Union>i. A i \<inter> s)" by blast
   have inc: "increasing M f"
     by (metis additive_increasing ca countably_additive_additive posf)
-  have sums: "psuminf (\<lambda>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
-    proof (rule countably_additiveD [OF ca])
+  have sums: "(\<Sum>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
+    proof (rule ca[unfolded countably_additive_def, rule_format])
       show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
         by blast
       show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
@@ -509,12 +530,14 @@
       show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
         by (metis UN_extend_simps(4) s seq)
     qed
-  hence "f s = psuminf (\<lambda>i. f (A i \<inter> s))"
+  hence "f s = (\<Sum>i. f (A i \<inter> s))"
     using seq [symmetric] by (simp add: sums_iff)
-  also have "... \<le> psuminf (f \<circ> A)"
-    proof (rule psuminf_le)
-      fix n show "f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
+  also have "... \<le> (\<Sum>i. f (A i))"
+    proof (rule suminf_le_pos)
+      fix n show "f (A n \<inter> s) \<le> f (A n)" using A s
         by (force intro: increasingD [OF inc])
+      fix N have "A N \<inter> s \<in> sets M"  using A s by auto
+      then show "0 \<le> f (A N \<inter> s)" using posf unfolding positive_def by auto
     qed
   also have "... = z" by (rule si)
   finally show "f s \<le> z" .
@@ -525,18 +548,40 @@
     by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl])
 qed
 
+lemma measure_set_pos:
+  assumes posf: "positive M f" "r \<in> measure_set M f X"
+  shows "0 \<le> r"
+proof -
+  obtain A where "range A \<subseteq> sets M" and r: "r = (\<Sum>i. f (A i))"
+    using `r \<in> measure_set M f X` unfolding measure_set_def by auto
+  then show "0 \<le> r" using posf unfolding r positive_def
+    by (intro suminf_0_le) auto
+qed
+
+lemma inf_measure_pos:
+  assumes posf: "positive M f"
+  shows "0 \<le> Inf (measure_set M f X)"
+proof (rule complete_lattice_class.Inf_greatest)
+  fix r assume "r \<in> measure_set M f X" with posf show "0 \<le> r"
+    by (rule measure_set_pos)
+qed
+
 lemma inf_measure_empty:
-  assumes posf: "positive M f" "{} \<in> sets M"
+  assumes posf: "positive M f" and "{} \<in> sets M"
   shows "Inf (measure_set M f {}) = 0"
 proof (rule antisym)
   show "Inf (measure_set M f {}) \<le> 0"
     by (metis complete_lattice_class.Inf_lower `{} \<in> sets M`
               inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
-qed simp
+qed (rule inf_measure_pos[OF posf])
 
 lemma (in algebra) inf_measure_positive:
-  "positive M f \<Longrightarrow> positive M (\<lambda>x. Inf (measure_set M f x))"
-  by (simp add: positive_def inf_measure_empty)
+  assumes p: "positive M f" and "{} \<in> sets M"
+  shows "positive M (\<lambda>x. Inf (measure_set M f x))"
+proof (unfold positive_def, intro conjI ballI)
+  show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
+  fix A assume "A \<in> sets M"
+qed (rule inf_measure_pos[OF p])
 
 lemma (in algebra) inf_measure_increasing:
   assumes posf: "positive M f"
@@ -548,25 +593,25 @@
 apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
 done
 
-
 lemma (in algebra) inf_measure_le:
   assumes posf: "positive M f" and inc: "increasing M f"
-      and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> psuminf (f \<circ> A) = r}"
+      and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
   shows "Inf (measure_set M f s) \<le> x"
 proof -
-  from x
   obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
-             and xeq: "psuminf (f \<circ> A) = x"
-    by auto
+             and xeq: "(\<Sum>i. f (A i)) = x"
+    using x by auto
   have dA: "range (disjointed A) \<subseteq> sets M"
     by (metis A range_disjointed_sets)
-  have "\<forall>n.(f o disjointed A) n \<le> (f \<circ> A) n" unfolding comp_def
+  have "\<forall>n. f (disjointed A n) \<le> f (A n)"
     by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
-  hence sda: "psuminf (f o disjointed A) \<le> psuminf (f \<circ> A)"
-    by (blast intro: psuminf_le)
-  hence ley: "psuminf (f o disjointed A) \<le> x"
+  moreover have "\<forall>i. 0 \<le> f (disjointed A i)"
+    using posf dA unfolding positive_def by auto
+  ultimately have sda: "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
+    by (blast intro!: suminf_le_pos)
+  hence ley: "(\<Sum>i. f (disjointed A i)) \<le> x"
     by (metis xeq)
-  hence y: "psuminf (f o disjointed A) \<in> measure_set M f s"
+  hence y: "(\<Sum>i. f (disjointed A i)) \<in> measure_set M f s"
     apply (auto simp add: measure_set_def)
     apply (rule_tac x="disjointed A" in exI)
     apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def)
@@ -576,13 +621,16 @@
 qed
 
 lemma (in algebra) inf_measure_close:
+  fixes e :: extreal
   assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
   shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
-               psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
-proof (cases "Inf (measure_set M f s) = \<omega>")
+               (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
+proof (cases "Inf (measure_set M f s) = \<infinity>")
   case False
+  then have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
+    using inf_measure_pos[OF posf, of s] by auto
   obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
-    using Inf_close[OF False e] by auto
+    using Inf_extreal_close[OF fin e] by auto
   thus ?thesis
     by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
 next
@@ -600,9 +648,8 @@
   shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
                   (\<lambda>x. Inf (measure_set M f x))"
   unfolding countably_subadditive_def o_def
-proof (safe, simp, rule pextreal_le_epsilon)
-  fix A :: "nat \<Rightarrow> 'a set" and e :: pextreal
-
+proof (safe, simp, rule extreal_le_epsilon, safe)
+  fix A :: "nat \<Rightarrow> 'a set" and e :: extreal
   let "?outer n" = "Inf (measure_set M f (A n))"
   assume A: "range A \<subseteq> Pow (space M)"
      and disj: "disjoint_family A"
@@ -610,21 +657,27 @@
      and e: "0 < e"
   hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
                    A n \<subseteq> (\<Union>i. BB n i) \<and>
-                   psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
-    apply (safe intro!: choice inf_measure_close [of f, OF posf _])
-    using e sb by (cases e, auto simp add: not_le mult_pos_pos)
+                   (\<Sum>i. f (BB n i)) \<le> ?outer n + e * (1/2)^(Suc n)"
+    apply (safe intro!: choice inf_measure_close [of f, OF posf])
+    using e sb by (cases e) (auto simp add: not_le mult_pos_pos one_extreal_def)
   then obtain BB
     where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
       and disjBB: "\<And>n. disjoint_family (BB n)"
       and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
-      and BBle: "\<And>n. psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
+      and BBle: "\<And>n. (\<Sum>i. f (BB n i)) \<le> ?outer n + e * (1/2)^(Suc n)"
     by auto blast
-  have sll: "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> psuminf ?outer + e"
+  have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> suminf ?outer + e"
     proof -
-      have "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> (\<Sum>\<^isub>\<infinity> n. ?outer n + e*(1/2) ^ Suc n)"
-        by (rule psuminf_le[OF BBle])
-      also have "... = psuminf ?outer + e"
-        using psuminf_half_series by simp
+      have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
+        using suminf_half_series_extreal e
+        by (simp add: extreal_zero_le_0_iff zero_le_divide_extreal suminf_cmult_extreal)
+      have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto
+      then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le)
+      then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer n + e*(1/2) ^ Suc n)"
+        by (rule suminf_le_pos[OF BBle])
+      also have "... = suminf ?outer + e"
+        using sum_eq_1 inf_measure_pos[OF posf] e
+        by (subst suminf_add_extreal) (auto simp add: extreal_zero_le_0_iff)
       finally show ?thesis .
     qed
   def C \<equiv> "(split BB) o prod_decode"
@@ -644,23 +697,25 @@
     qed
   have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
     by (rule ext)  (auto simp add: C_def)
-  moreover have "psuminf ... = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" using BBle
-    by (force intro!: psuminf_2dimen simp: o_def)
-  ultimately have Csums: "psuminf (f \<circ> C) = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" by simp
-  have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))"
+  moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle
+    using BB posf[unfolded positive_def]
+    by (force intro!: suminf_extreal_2dimen simp: o_def)
+  ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def)
+  have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>n. \<Sum>i. f (BB n i))"
     apply (rule inf_measure_le [OF posf(1) inc], auto)
     apply (rule_tac x="C" in exI)
     apply (auto simp add: C sbC Csums)
     done
-  also have "... \<le> (\<Sum>\<^isub>\<infinity>n. Inf (measure_set M f (A n))) + e" using sll
+  also have "... \<le> (\<Sum>n. Inf (measure_set M f (A n))) + e" using sll
     by blast
-  finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> psuminf ?outer + e" .
+  finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> suminf ?outer + e" .
 qed
 
 lemma (in algebra) inf_measure_outer:
   "\<lbrakk> positive M f ; increasing M f \<rbrakk>
    \<Longrightarrow> outer_measure_space \<lparr> space = space M, sets = Pow (space M) \<rparr>
                           (\<lambda>x. Inf (measure_set M f x))"
+  using inf_measure_pos[of M f]
   by (simp add: outer_measure_space_def inf_measure_empty
                 inf_measure_increasing inf_measure_countably_subadditive positive_def)
 
@@ -680,13 +735,13 @@
     by blast
   have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
         \<le> Inf (measure_set M f s)"
-    proof (rule pextreal_le_epsilon)
-      fix e :: pextreal
+    proof (rule extreal_le_epsilon, intro allI impI)
+      fix e :: extreal
       assume e: "0 < e"
       from inf_measure_close [of f, OF posf e s]
       obtain A where A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
                  and sUN: "s \<subseteq> (\<Union>i. A i)"
-                 and l: "psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
+                 and l: "(\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
         by auto
       have [simp]: "!!x. x \<in> sets M \<Longrightarrow>
                       (f o (\<lambda>z. z \<inter> (space M - x)) o A) = (f o (\<lambda>z. z - x) o A)"
@@ -698,9 +753,9 @@
         assume u: "u \<in> sets M"
         have [simp]: "\<And>n. f (A n \<inter> u) \<le> f (A n)"
           by (simp add: increasingD [OF inc] u Int range_subsetD [OF A])
-        have 2: "Inf (measure_set M f (s \<inter> u)) \<le> psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A)"
+        have 2: "Inf (measure_set M f (s \<inter> u)) \<le> (\<Sum>i. f (A i \<inter> u))"
           proof (rule complete_lattice_class.Inf_lower)
-            show "psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A) \<in> measure_set M f (s \<inter> u)"
+            show "(\<Sum>i. f (A i \<inter> u)) \<in> measure_set M f (s \<inter> u)"
               apply (simp add: measure_set_def)
               apply (rule_tac x="(\<lambda>z. z \<inter> u) o A" in exI)
               apply (auto simp add: disjoint_family_subset [OF disj] o_def)
@@ -709,15 +764,16 @@
               done
           qed
       } note lesum = this
-      have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> psuminf (f o (\<lambda>z. z\<inter>x) o A)"
+      have [simp]: "\<And>i. A i \<inter> (space M - x) = A i - x" using A sets_into_space by auto
+      have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> (\<Sum>i. f (A i \<inter> x))"
         and inf2: "Inf (measure_set M f (s \<inter> (space M - x)))
-                   \<le> psuminf (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
+                   \<le> (\<Sum>i. f (A i \<inter> (space M - x)))"
         by (metis Diff lesum top x)+
       hence "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
-           \<le> psuminf (f o (\<lambda>s. s\<inter>x) o A) + psuminf (f o (\<lambda>s. s-x) o A)"
-        by (simp add: x add_mono)
-      also have "... \<le> psuminf (f o A)"
-        by (simp add: x psuminf_add[symmetric] o_def)
+           \<le> (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))"
+        by (simp add: add_mono x)
+      also have "... \<le> (\<Sum>i. f (A i))" using posf[unfolded positive_def] A x
+        by (subst suminf_add_extreal[symmetric]) auto
       also have "... \<le> Inf (measure_set M f s) + e"
         by (rule l)
       finally show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
@@ -732,7 +788,7 @@
     also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
       apply (rule subadditiveD)
       apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow])
-      apply (simp add: positive_def inf_measure_empty[OF posf])
+      apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf])
       apply (rule inf_measure_countably_subadditive)
       using s by (auto intro!: posf inc)
     finally show ?thesis .
@@ -751,7 +807,7 @@
 
 theorem (in algebra) caratheodory:
   assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pextreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
             measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
 proof -
   have inc: "increasing M f"
--- a/src/HOL/Probability/Complete_Measure.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,7 +1,6 @@
-(*  Title:      HOL/Probability/Complete_Measure.thy
+(*  Title:      Complete_Measure.thy
     Author:     Robert Himmelmann, Johannes Hoelzl, TU Muenchen
 *)
-
 theory Complete_Measure
 imports Product_Measure
 begin
@@ -177,7 +176,8 @@
 proof -
   show "measure_space completion"
   proof
-    show "measure completion {} = 0" by (auto simp: completion_def)
+    show "positive completion (measure completion)"
+      by (auto simp: completion_def positive_def)
   next
     show "countably_additive completion (measure completion)"
     proof (intro countably_additiveI)
@@ -189,9 +189,9 @@
           using A by (subst (1 2) main_part_null_part_Un) auto
         then show "main_part (A n) \<inter> main_part (A m) = {}" by auto
       qed
-      then have "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
+      then have "(\<Sum>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
         unfolding completion_def using A by (auto intro!: measure_countably_additive)
-      then show "(\<Sum>\<^isub>\<infinity>n. measure completion (A n)) = measure completion (UNION UNIV A)"
+      then show "(\<Sum>n. measure completion (A n)) = measure completion (UNION UNIV A)"
         by (simp add: completion_def \<mu>_main_part_UN[OF A(1)])
     qed
   qed
@@ -251,30 +251,52 @@
   qed
 qed
 
-lemma (in completeable_measure_space) completion_ex_borel_measurable:
-  fixes g :: "'a \<Rightarrow> pextreal"
-  assumes g: "g \<in> borel_measurable completion"
+lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
+  fixes g :: "'a \<Rightarrow> extreal"
+  assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
   shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
 proof -
-  from g[THEN completion.borel_measurable_implies_simple_function_sequence]
-  obtain f where "\<And>i. simple_function completion (f i)" "f \<up> g" by auto
-  then have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)"
-    using completion_ex_simple_function by auto
+  from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this
+  from this(1)[THEN completion_ex_simple_function]
+  have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)" ..
   from this[THEN choice] obtain f' where
     sf: "\<And>i. simple_function M (f' i)" and
     AE: "\<forall>i. AE x. f i x = f' i x" by auto
   show ?thesis
   proof (intro bexI)
-    from AE[unfolded all_AE_countable]
+    from AE[unfolded AE_all_countable[symmetric]]
     show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x")
     proof (elim AE_mp, safe intro!: AE_I2)
       fix x assume eq: "\<forall>i. f i x = f' i x"
-      moreover have "g = SUPR UNIV f" using `f \<up> g` unfolding isoton_def by simp
-      ultimately show "g x = ?f x" by (simp add: SUPR_apply)
+      moreover have "g x = (SUP i. f i x)"
+        unfolding f using `0 \<le> g x` by (auto split: split_max)
+      ultimately show "g x = ?f x" by auto
     qed
     show "?f \<in> borel_measurable M"
       using sf by (auto intro: borel_measurable_simple_function)
   qed
 qed
 
+lemma (in completeable_measure_space) completion_ex_borel_measurable:
+  fixes g :: "'a \<Rightarrow> extreal"
+  assumes g: "g \<in> borel_measurable completion"
+  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
+proof -
+  have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
+  from completion_ex_borel_measurable_pos[OF this] guess g_pos ..
+  moreover
+  have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
+  from completion_ex_borel_measurable_pos[OF this] guess g_neg ..
+  ultimately
+  show ?thesis
+  proof (safe intro!: bexI[of _ "\<lambda>x. g_pos x - g_neg x"])
+    show "AE x. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
+    proof (intro AE_I2 impI)
+      fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x"
+      show "g x = g_pos x - g_neg x" unfolding g[symmetric]
+        by (cases "g x") (auto split: split_max)
+    qed
+  qed auto
+qed
+
 end
--- a/src/HOL/Probability/Information.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Information.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -2,9 +2,12 @@
 imports
   Probability_Space
   "~~/src/HOL/Library/Convex"
-  Lebesgue_Measure
 begin
 
+lemma (in prob_space) not_zero_less_distribution[simp]:
+  "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0"
+  using distribution_positive[of X A] by arith
+
 lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
   by (subst log_le_cancel_iff) auto
 
@@ -238,7 +241,7 @@
   have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
   show "(\<Sum>x \<in> space M. log b (real (RN_deriv M \<nu> x)) * real (\<nu> {x})) = ?sum"
     using RN_deriv_finite_measure[OF ms ac]
-    by (auto intro!: setsum_cong simp: field_simps real_of_pextreal_mult[symmetric])
+    by (auto intro!: setsum_cong simp: field_simps)
 qed
 
 lemma (in finite_prob_space) KL_divergence_positive_finite:
@@ -254,7 +257,8 @@
   proof (subst KL_divergence_eq_finite[OF ms ac], safe intro!: log_setsum_divide not_empty)
     show "finite (space M)" using finite_space by simp
     show "1 < b" by fact
-    show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1" using v.finite_sum_over_space_eq_1 by simp
+    show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1"
+      using v.finite_sum_over_space_eq_1 by (simp add: v.\<mu>'_def)
 
     fix x assume "x \<in> space M"
     then have x: "{x} \<in> sets M" unfolding sets_eq_Pow by auto
@@ -262,17 +266,19 @@
       then have "\<nu> {x} \<noteq> 0" by auto
       then have "\<mu> {x} \<noteq> 0"
         using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto
-      thus "0 < prob {x}" using finite_measure[of "{x}"] x by auto }
-  qed auto
-  thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by simp
+      thus "0 < real (\<mu> {x})" using real_measure[OF x] by auto }
+    show "0 \<le> real (\<mu> {x})" "0 \<le> real (\<nu> {x})"
+      using real_measure[OF x] v.real_measure[of "{x}"] x by auto
+  qed
+  thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by (simp add: \<mu>'_def)
 qed
 
 subsection {* Mutual Information *}
 
 definition (in prob_space)
   "mutual_information b S T X Y =
-    KL_divergence b (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>)
-      (joint_distribution X Y)"
+    KL_divergence b (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
+      (extreal\<circ>joint_distribution X Y)"
 
 definition (in prob_space)
   "entropy b s X = mutual_information b s s X X"
@@ -280,38 +286,33 @@
 abbreviation (in information_space)
   mutual_information_Pow ("\<I>'(_ ; _')") where
   "\<I>(X ; Y) \<equiv> mutual_information b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr> X Y"
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
 
 lemma (in prob_space) finite_variables_absolutely_continuous:
   assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
   shows "measure_space.absolutely_continuous
-    (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>)
-    (joint_distribution X Y)"
+    (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
+    (extreal\<circ>joint_distribution X Y)"
 proof -
-  interpret X: finite_prob_space "S\<lparr>measure := distribution X\<rparr>"
+  interpret X: finite_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
     using X by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "T\<lparr>measure := distribution Y\<rparr>"
+  interpret Y: finite_prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
     using Y by (rule distribution_finite_prob_space)
   interpret XY: pair_finite_prob_space
-    "S\<lparr>measure := distribution X\<rparr>" "T\<lparr> measure := distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr> measure := joint_distribution X Y\<rparr>"
+    "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr> measure := extreal\<circ>distribution Y\<rparr>" by default
+  interpret P: finite_prob_space "XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>"
     using assms by (auto intro!: joint_distribution_finite_prob_space)
   note rv = assms[THEN finite_random_variableD]
-  show "XY.absolutely_continuous (joint_distribution X Y)"
+  show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
   proof (rule XY.absolutely_continuousI)
-    show "finite_measure_space (XY.P\<lparr> measure := joint_distribution X Y\<rparr>)" by default
+    show "finite_measure_space (XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
     fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
-    then obtain a b where "(a, b) = x" and "a \<in> space S" "b \<in> space T"
-      and distr: "distribution X {a} * distribution Y {b} = 0"
+    then obtain a b where "x = (a, b)"
+      and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
       by (cases x) (auto simp: space_pair_measure)
-    with X.sets_eq_Pow Y.sets_eq_Pow
-      joint_distribution_Times_le_fst[OF rv, of "{a}" "{b}"]
-      joint_distribution_Times_le_snd[OF rv, of "{a}" "{b}"]
-    have "joint_distribution X Y {x} \<le> distribution Y {b}"
-         "joint_distribution X Y {x} \<le> distribution X {a}"
-      by (auto simp del: X.sets_eq_Pow Y.sets_eq_Pow)
-    with distr show "joint_distribution X Y {x} = 0" by auto
+    with finite_distribution_order(5,6)[OF X Y]
+    show "(extreal \<circ> joint_distribution X Y) {x} = 0" by auto
   qed
 qed
 
@@ -320,28 +321,28 @@
   assumes MY: "finite_random_variable MY Y"
   shows mutual_information_generic_eq:
     "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
-      real (joint_distribution X Y {(x,y)}) *
-      log b (real (joint_distribution X Y {(x,y)}) /
-      (real (distribution X {x}) * real (distribution Y {y}))))"
+      joint_distribution X Y {(x,y)} *
+      log b (joint_distribution X Y {(x,y)} /
+      (distribution X {x} * distribution Y {y})))"
     (is ?sum)
   and mutual_information_positive_generic:
      "0 \<le> mutual_information b MX MY X Y" (is ?positive)
 proof -
-  interpret X: finite_prob_space "MX\<lparr>measure := distribution X\<rparr>"
+  interpret X: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
     using MX by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "MY\<lparr>measure := distribution Y\<rparr>"
+  interpret Y: finite_prob_space "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
     using MY by (rule distribution_finite_prob_space)
-  interpret XY: pair_finite_prob_space "MX\<lparr>measure := distribution X\<rparr>" "MY\<lparr>measure := distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr>measure := joint_distribution X Y\<rparr>"
+  interpret XY: pair_finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>" "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+  interpret P: finite_prob_space "XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>"
     using assms by (auto intro!: joint_distribution_finite_prob_space)
 
-  have P_ms: "finite_measure_space (XY.P\<lparr>measure :=joint_distribution X Y\<rparr>)" by default
-  have P_ps: "finite_prob_space (XY.P\<lparr>measure := joint_distribution X Y\<rparr>)" by default
+  have P_ms: "finite_measure_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
+  have P_ps: "finite_prob_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
 
   show ?sum
     unfolding Let_def mutual_information_def
     by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]])
-       (auto simp add: space_pair_measure setsum_cartesian_product' real_of_pextreal_mult[symmetric])
+       (auto simp add: space_pair_measure setsum_cartesian_product')
 
   show ?positive
     using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
@@ -351,10 +352,10 @@
 lemma (in information_space) mutual_information_commute_generic:
   assumes X: "random_variable S X" and Y: "random_variable T Y"
   assumes ac: "measure_space.absolutely_continuous
-    (S\<lparr>measure := distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := distribution Y\<rparr>) (joint_distribution X Y)"
+    (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>) (extreal\<circ>joint_distribution X Y)"
   shows "mutual_information b S T X Y = mutual_information b T S Y X"
 proof -
-  let ?S = "S\<lparr>measure := distribution X\<rparr>" and ?T = "T\<lparr>measure := distribution Y\<rparr>"
+  let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
   interpret S: prob_space ?S using X by (rule distribution_prob_space)
   interpret T: prob_space ?T using Y by (rule distribution_prob_space)
   interpret P: pair_prob_space ?S ?T ..
@@ -363,13 +364,13 @@
     unfolding mutual_information_def
   proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1])
     show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
-      (P.P \<lparr> measure := joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := joint_distribution Y X\<rparr>)"
+      (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
       using X Y unfolding measurable_def
       unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable
-      by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>])
-    have "prob_space (P.P\<lparr> measure := joint_distribution X Y\<rparr>)"
+      by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>'])
+    have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
       using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
-    then show "measure_space (P.P\<lparr> measure := joint_distribution X Y\<rparr>)"
+    then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
       unfolding prob_space_def by simp
   qed auto
 qed
@@ -389,8 +390,8 @@
 lemma (in information_space) mutual_information_eq:
   assumes "simple_function M X" "simple_function M Y"
   shows "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
-    real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) /
-                                                   (real (distribution X {x}) * real (distribution Y {y}))))"
+    distribution (\<lambda>x. (X x, Y x)) {(x,y)} * log b (distribution (\<lambda>x. (X x, Y x)) {(x,y)} /
+                                                   (distribution X {x} * distribution Y {y})))"
   using assms by (simp add: mutual_information_generic_eq)
 
 lemma (in information_space) mutual_information_generic_cong:
@@ -416,22 +417,27 @@
 
 abbreviation (in information_space)
   entropy_Pow ("\<H>'(_')") where
-  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr> X"
+  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr> X"
 
 lemma (in information_space) entropy_generic_eq:
+  fixes X :: "'a \<Rightarrow> 'c"
   assumes MX: "finite_random_variable MX X"
-  shows "entropy b MX X = -(\<Sum> x \<in> space MX. real (distribution X {x}) * log b (real (distribution X {x})))"
+  shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
 proof -
-  interpret MX: finite_prob_space "MX\<lparr>measure := distribution X\<rparr>"
+  interpret MX: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
     using MX by (rule distribution_finite_prob_space)
-  let "?X x" = "real (distribution X {x})"
-  let "?XX x y" = "real (joint_distribution X X {(x, y)})"
-  { fix x y
-    have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
+  let "?X x" = "distribution X {x}"
+  let "?XX x y" = "joint_distribution X X {(x, y)}"
+
+  { fix x y :: 'c
+    { assume "x \<noteq> y"
+      then have "(\<lambda>x. (X x, X x)) -` {(x,y)} \<inter> space M = {}" by auto
+      then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) }
     then have "?XX x y * log b (?XX x y / (?X x * ?X y)) =
         (if x = y then - ?X y * log b (?X y) else 0)"
-      unfolding distribution_def by (auto simp: log_simps zero_less_mult_iff) }
+      by (auto simp: log_simps zero_less_mult_iff) }
   note remove_XX = this
+
   show ?thesis
     unfolding entropy_def mutual_information_generic_eq[OF MX MX]
     unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX
@@ -440,7 +446,7 @@
 
 lemma (in information_space) entropy_eq:
   assumes "simple_function M X"
-  shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
+  shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))"
   using assms by (simp add: entropy_generic_eq)
 
 lemma (in information_space) entropy_positive:
@@ -448,63 +454,77 @@
   unfolding entropy_def by (simp add: mutual_information_positive)
 
 lemma (in information_space) entropy_certainty_eq_0:
-  assumes "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
+  assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
   shows "\<H>(X) = 0"
 proof -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
+  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal\<circ>distribution X\<rparr>"
   note simple_function_imp_finite_random_variable[OF `simple_function M X`]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = distribution X \<rparr>"]
+  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
   interpret X: finite_prob_space ?X by simp
   have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
     using X.measure_compl[of "{x}"] assms by auto
   also have "\<dots> = 0" using X.prob_space assms by auto
   finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
-  { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
-    hence "{y} \<subseteq> X ` space M - {x}" by auto
-    from X.measure_mono[OF this] X0 asm
-    have "distribution X {y} = 0" by auto }
-  hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
-    using assms by auto
+  { fix y assume *: "y \<in> X ` space M"
+    { assume asm: "y \<noteq> x"
+      with * have "{y} \<subseteq> X ` space M - {x}" by auto
+      from X.measure_mono[OF this] X0 asm *
+      have "distribution X {y} = 0"  by (auto intro: antisym) }
+    then have "distribution X {y} = (if x = y then 1 else 0)"
+      using assms by auto }
+  note fi = this
   have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
   show ?thesis unfolding entropy_eq[OF `simple_function M X`] by (auto simp: y fi)
 qed
 
 lemma (in information_space) entropy_le_card_not_0:
-  assumes "simple_function M X"
-  shows "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
+  assumes X: "simple_function M X"
+  shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}))"
 proof -
-  let "?d x" = "distribution X {x}"
-  let "?p x" = "real (?d x)"
+  let "?p x" = "distribution X {x}"
   have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
-    by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function M X`] setsum_negf[symmetric] log_simps not_less)
+    unfolding entropy_eq[OF X] setsum_negf[symmetric]
+    by (auto intro!: setsum_cong simp: log_simps)
   also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
-    apply (rule log_setsum')
-    using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution
-    by (auto simp: simple_function_def)
-  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
-    using distribution_finite[OF `simple_function M X`[THEN simple_function_imp_random_variable], simplified]
-    by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) (auto simp: real_of_pextreal_eq_0)
+    using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X]
+    by (intro log_setsum') (auto simp: simple_function_def)
+  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?p x \<noteq> 0 then 1 else 0)"
+    by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) auto
   finally show ?thesis
     using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
 qed
 
+lemma (in prob_space) measure'_translate:
+  assumes X: "random_variable S X" and A: "A \<in> sets S"
+  shows "finite_measure.\<mu>' (S\<lparr> measure := extreal\<circ>distribution X \<rparr>) A = distribution X A"
+proof -
+  interpret S: prob_space "S\<lparr> measure := extreal\<circ>distribution X \<rparr>"
+    using distribution_prob_space[OF X] .
+  from A show "S.\<mu>' A = distribution X A"
+    unfolding S.\<mu>'_def by (simp add: distribution_def_raw \<mu>'_def)
+qed
+
 lemma (in information_space) entropy_uniform_max:
-  assumes "simple_function M X"
+  assumes X: "simple_function M X"
   assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
   shows "\<H>(X) = log b (real (card (X ` space M)))"
 proof -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
-  note simple_function_imp_finite_random_variable[OF `simple_function M X`]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = distribution X \<rparr>"]
+  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := extreal\<circ>distribution X\<rparr>"
+  note frv = simple_function_imp_finite_random_variable[OF X]
+  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
   interpret X: finite_prob_space ?X by simp
+  note rv = finite_random_variableD[OF frv]
   have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
     using `simple_function M X` not_empty by (auto simp: simple_function_def)
-  { fix x assume "x \<in> X ` space M"
-    hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
-    proof (rule X.uniform_prob[simplified])
-      fix x y assume "x \<in> X`space M" "y \<in> X`space M"
-      from assms(2)[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
-    qed }
+  { fix x assume "x \<in> space ?X"
+    moreover then have "X.\<mu>' {x} = 1 / card (space ?X)"
+    proof (rule X.uniform_prob)
+      fix x y assume "x \<in> space ?X" "y \<in> space ?X"
+      with assms(2)[of x y] show "X.\<mu>' {x} = X.\<mu>' {y}"
+        by (subst (1 2) measure'_translate[OF rv]) auto
+    qed
+    ultimately have "distribution X {x} = 1 / card (space ?X)"
+      by (subst (asm) measure'_translate[OF rv]) auto }
   thus ?thesis
     using not_empty X.finite_space b_gt_1 card_gt0
     by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps)
@@ -552,8 +572,7 @@
 lemma (in information_space) entropy_eq_cartesian_product:
   assumes "simple_function M X" "simple_function M Y"
   shows "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
-    real (joint_distribution X Y {(x,y)}) *
-    log b (real (joint_distribution X Y {(x,y)})))"
+    joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))"
 proof -
   have sf: "simple_function M (\<lambda>x. (X x, Y x))"
     using assms by (auto intro: simple_function_Pair)
@@ -576,9 +595,9 @@
 abbreviation (in information_space)
   conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
   "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr>
-    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = distribution Z \<rparr>
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr>
+    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = extreal\<circ>distribution Z \<rparr>
     X Y Z"
 
 lemma (in information_space) conditional_mutual_information_generic_eq:
@@ -586,58 +605,44 @@
     and MY: "finite_random_variable MY Y"
     and MZ: "finite_random_variable MZ Z"
   shows "conditional_mutual_information b MX MY MZ X Y Z = (\<Sum>(x, y, z) \<in> space MX \<times> space MY \<times> space MZ.
-             real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
-             log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
-    (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
-  (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z)))")
+             distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
+             log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
+    (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
+  (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))")
 proof -
-  let ?YZ = "\<lambda>y z. real (joint_distribution Y Z {(y, z)})"
-  let ?X = "\<lambda>x. real (distribution X {x})"
-  let ?Z = "\<lambda>z. real (distribution Z {z})"
-
-  txt {* This proof is actually quiet easy, however we need to show that the
-    distributions are finite and the joint distributions are zero when one of
-    the variables distribution is also zero. *}
-
+  let ?X = "\<lambda>x. distribution X {x}"
   note finite_var = MX MY MZ
-  note random_var = finite_var[THEN finite_random_variableD]
-
-  note space_simps = space_pair_measure space_sigma algebra.simps
-
   note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
+  note XYZ = finite_random_variable_pairI[OF MX YZ]
   note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
   note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
   note YZX = finite_random_variable_pairI[OF finite_var(2) ZX]
   note order1 =
-    finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
-    finite_distribution_order(5,6)[OF finite_var(1,3), simplified space_simps]
+    finite_distribution_order(5,6)[OF finite_var(1) YZ]
+    finite_distribution_order(5,6)[OF finite_var(1,3)]
 
+  note random_var = finite_var[THEN finite_random_variableD]
   note finite = finite_var(1) YZ finite_var(3) XZ YZX
-  note finite[THEN finite_distribution_finite, simplified space_simps, simp]
 
   have order2: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
           \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
     unfolding joint_distribution_commute_singleton[of X]
     unfolding joint_distribution_assoc_singleton[symmetric]
     using finite_distribution_order(6)[OF finite_var(2) ZX]
-    by (auto simp: space_simps)
+    by auto
 
-  have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z))) =
+  have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z)))) =
     (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))"
     (is "(\<Sum>(x, y, z)\<in>?S. ?L x y z) = (\<Sum>(x, y, z)\<in>?S. ?R x y z)")
   proof (safe intro!: setsum_cong)
     fix x y z assume space: "x \<in> space MX" "y \<in> space MY" "z \<in> space MZ"
-    then have *: "?XYZ x y z / (?XZ x z * ?YZdZ y z) =
-      (?XYZ x y z / (?X x * ?YZ y z)) / (?XZ x z / (?X x * ?Z z))"
-      using order1(3)
-      by (auto simp: real_of_pextreal_mult[symmetric] real_of_pextreal_eq_0)
     show "?L x y z = ?R x y z"
     proof cases
       assume "?XYZ x y z \<noteq> 0"
-      with space b_gt_1 order1 order2 show ?thesis unfolding *
-        by (subst log_divide)
-           (auto simp: zero_less_divide_iff zero_less_real_of_pextreal
-                       real_of_pextreal_eq_0 zero_less_mult_iff)
+      with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z"
+        using order1 order2 by (auto simp: less_le)
+      with b_gt_1 show ?thesis
+        by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff)
     qed simp
   qed
   also have "\<dots> = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
@@ -649,8 +654,8 @@
               setsum_left_distrib[symmetric]
     unfolding joint_distribution_commute_singleton[of X]
     unfolding joint_distribution_assoc_singleton[symmetric]
-    using setsum_real_joint_distribution_singleton[OF finite_var(2) ZX, unfolded space_simps]
-    by (intro setsum_cong refl) simp
+    using setsum_joint_distribution_singleton[OF finite_var(2) ZX]
+    by (intro setsum_cong refl) (simp add: space_pair_measure)
   also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
              (\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) =
              conditional_mutual_information b MX MY MZ X Y Z"
@@ -664,11 +669,11 @@
 lemma (in information_space) conditional_mutual_information_eq:
   assumes "simple_function M X" "simple_function M Y" "simple_function M Z"
   shows "\<I>(X;Y|Z) = (\<Sum>(x, y, z) \<in> X`space M \<times> Y`space M \<times> Z`space M.
-             real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
-             log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
-    (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
-  using conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
-  by simp
+             distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
+             log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
+    (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))"
+  by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
+     simp
 
 lemma (in information_space) conditional_mutual_information_eq_mutual_information:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
@@ -683,10 +688,10 @@
 qed
 
 lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
-  unfolding distribution_def using measure_space_1 by auto
+  unfolding distribution_def using prob_space by auto
 
 lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\<lambda>x. (X x, ())) {(a, ())} = distribution X {a}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
 
 lemma (in prob_space) setsum_distribution:
   assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
@@ -695,12 +700,13 @@
 
 lemma (in prob_space) setsum_real_distribution:
   fixes MX :: "('c, 'd) measure_space_scheme"
-  assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. real (distribution X {a})) = 1"
-  using setsum_real_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV, measure = undefined \<rparr>" "\<lambda>x. ()" "{()}"]
-  using sigma_algebra_Pow[of "UNIV::unit set" "\<lparr> measure = undefined \<rparr>"] by simp
+  assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
+  using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV, measure = undefined \<rparr>" "\<lambda>x. ()" "{()}"]
+  using sigma_algebra_Pow[of "UNIV::unit set" "\<lparr> measure = undefined \<rparr>"]
+  by auto
 
 lemma (in information_space) conditional_mutual_information_generic_positive:
-  assumes "finite_random_variable MX X" and "finite_random_variable MY Y" and "finite_random_variable MZ Z"
+  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "finite_random_variable MZ Z"
   shows "0 \<le> conditional_mutual_information b MX MY MZ X Y Z"
 proof (cases "space MX \<times> space MY \<times> space MZ = {}")
   case True show ?thesis
@@ -708,43 +714,35 @@
     by simp
 next
   case False
-  let "?dXYZ A" = "real (distribution (\<lambda>x. (X x, Y x, Z x)) A)"
-  let "?dXZ A" = "real (joint_distribution X Z A)"
-  let "?dYZ A" = "real (joint_distribution Y Z A)"
-  let "?dX A" = "real (distribution X A)"
-  let "?dZ A" = "real (distribution Z A)"
+  let ?dXYZ = "distribution (\<lambda>x. (X x, Y x, Z x))"
+  let ?dXZ = "joint_distribution X Z"
+  let ?dYZ = "joint_distribution Y Z"
+  let ?dX = "distribution X"
+  let ?dZ = "distribution Z"
   let ?M = "space MX \<times> space MY \<times> space MZ"
 
-  have split_beta: "\<And>f. split f = (\<lambda>x. f (fst x) (snd x))" by (simp add: fun_eq_iff)
-
-  note space_simps = space_pair_measure space_sigma algebra.simps
-
-  note finite_var = assms
-  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
-  note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
-  note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
-  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
-  note XYZ = finite_random_variable_pairI[OF finite_var(1) YZ]
-  note finite = finite_var(3) YZ XZ XYZ
-  note finite = finite[THEN finite_distribution_finite, simplified space_simps]
-
+  note YZ = finite_random_variable_pairI[OF Y Z]
+  note XZ = finite_random_variable_pairI[OF X Z]
+  note ZX = finite_random_variable_pairI[OF Z X]
+  note YZ = finite_random_variable_pairI[OF Y Z]
+  note XYZ = finite_random_variable_pairI[OF X YZ]
+  note finite = Z YZ XZ XYZ
   have order: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
           \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
     unfolding joint_distribution_commute_singleton[of X]
     unfolding joint_distribution_assoc_singleton[symmetric]
-    using finite_distribution_order(6)[OF finite_var(2) ZX]
-    by (auto simp: space_simps)
+    using finite_distribution_order(6)[OF Y ZX]
+    by auto
 
   note order = order
-    finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
-    finite_distribution_order(5,6)[OF finite_var(2,3), simplified space_simps]
+    finite_distribution_order(5,6)[OF X YZ]
+    finite_distribution_order(5,6)[OF Y Z]
 
   have "- conditional_mutual_information b MX MY MZ X Y Z = - (\<Sum>(x, y, z) \<in> ?M. ?dXYZ {(x, y, z)} *
     log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))"
-    unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal
-    by (intro setsum_cong) (auto intro!: arg_cong[where f="log b"] simp: real_of_pextreal_mult[symmetric])
+    unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal by auto
   also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
-    unfolding split_beta
+    unfolding split_beta'
   proof (rule log_setsum_divide)
     show "?M \<noteq> {}" using False by simp
     show "1 < b" using b_gt_1 .
@@ -757,33 +755,31 @@
       unfolding setsum_commute[of _ "space MY"]
       unfolding setsum_commute[of _ "space MZ"]
       by (simp_all add: space_pair_measure
-        setsum_real_joint_distribution_singleton[OF `finite_random_variable MX X` YZ]
-        setsum_real_joint_distribution_singleton[OF `finite_random_variable MY Y` finite_var(3)]
-        setsum_real_distribution[OF `finite_random_variable MZ Z`])
+                        setsum_joint_distribution_singleton[OF X YZ]
+                        setsum_joint_distribution_singleton[OF Y Z]
+                        setsum_distribution[OF Z])
 
     fix x assume "x \<in> ?M"
     let ?x = "(fst x, fst (snd x), snd (snd x))"
 
-    show "0 \<le> ?dXYZ {?x}" using real_pextreal_nonneg .
-    show "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
-     by (simp add: real_pextreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg)
+    show "0 \<le> ?dXYZ {?x}"
+      "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
+     by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg)
 
     assume *: "0 < ?dXYZ {?x}"
-    with `x \<in> ?M` show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
-      using finite order
-      by (cases x)
-         (auto simp add: zero_less_real_of_pextreal zero_less_mult_iff zero_less_divide_iff)
+    with `x \<in> ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
+      by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le)
   qed
   also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>space MZ. ?dZ {z})"
     apply (simp add: setsum_cartesian_product')
     apply (subst setsum_commute)
     apply (subst (2) setsum_commute)
     by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric]
-                   setsum_real_joint_distribution_singleton[OF finite_var(1,3)]
-                   setsum_real_joint_distribution_singleton[OF finite_var(2,3)]
+                   setsum_joint_distribution_singleton[OF X Z]
+                   setsum_joint_distribution_singleton[OF Y Z]
           intro!: setsum_cong)
   also have "log b (\<Sum>z\<in>space MZ. ?dZ {z}) = 0"
-    unfolding setsum_real_distribution[OF finite_var(3)] by simp
+    unfolding setsum_real_distribution[OF Z] by simp
   finally show ?thesis by simp
 qed
 
@@ -800,57 +796,52 @@
 abbreviation (in information_space)
   conditional_entropy_Pow ("\<H>'(_ | _')") where
   "\<H>(X | Y) \<equiv> conditional_entropy b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \<rparr> X Y"
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
 
 lemma (in information_space) conditional_entropy_positive:
   "simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
   unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive)
 
-lemma (in measure_space) empty_measureI: "A = {} \<Longrightarrow> \<mu> A = 0" by simp
-
 lemma (in information_space) conditional_entropy_generic_eq:
   fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
   assumes MX: "finite_random_variable MX X"
   assumes MZ: "finite_random_variable MZ Z"
   shows "conditional_entropy b MX MZ X Z =
      - (\<Sum>(x, z)\<in>space MX \<times> space MZ.
-         real (joint_distribution X Z {(x, z)}) *
-         log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
+         joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
 proof -
   interpret MX: finite_sigma_algebra MX using MX by simp
   interpret MZ: finite_sigma_algebra MZ using MZ by simp
   let "?XXZ x y z" = "joint_distribution X (\<lambda>x. (X x, Z x)) {(x, y, z)}"
   let "?XZ x z" = "joint_distribution X Z {(x, z)}"
   let "?Z z" = "distribution Z {z}"
-  let "?f x y z" = "log b (real (?XXZ x y z) / (real (?XZ x z) * real (?XZ y z / ?Z z)))"
+  let "?f x y z" = "log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y z))"
   { fix x z have "?XXZ x x z = ?XZ x z"
-      unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) }
+      unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>']) }
   note this[simp]
   { fix x x' :: 'c and z assume "x' \<noteq> x"
     then have "?XXZ x x' z = 0"
-      by (auto simp: distribution_def intro!: arg_cong[where f=\<mu>] empty_measureI) }
+      by (auto simp: distribution_def empty_measure'[symmetric]
+               simp del: empty_measure' intro!: arg_cong[where f=\<mu>']) }
   note this[simp]
   { fix x x' z assume *: "x \<in> space MX" "z \<in> space MZ"
-    then have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z)
-      = (\<Sum>x'\<in>space MX. if x = x' then real (?XZ x z) * ?f x x z else 0)"
+    then have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z)
+      = (\<Sum>x'\<in>space MX. if x = x' then ?XZ x z * ?f x x z else 0)"
       by (auto intro!: setsum_cong)
-    also have "\<dots> = real (?XZ x z) * ?f x x z"
+    also have "\<dots> = ?XZ x z * ?f x x z"
       using `x \<in> space MX` by (simp add: setsum_cases[OF MX.finite_space])
-    also have "\<dots> = real (?XZ x z) * log b (real (?Z z) / real (?XZ x z))"
-      by (auto simp: real_of_pextreal_mult[symmetric])
-    also have "\<dots> = - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))"
-      using assms[THEN finite_distribution_finite]
+    also have "\<dots> = ?XZ x z * log b (?Z z / ?XZ x z)" by auto
+    also have "\<dots> = - ?XZ x z * log b (?XZ x z / ?Z z)"
       using finite_distribution_order(6)[OF MX MZ]
-      by (auto simp: log_simps field_simps zero_less_mult_iff zero_less_real_of_pextreal real_of_pextreal_eq_0)
-    finally have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z) =
-      - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" . }
+      by (auto simp: log_simps field_simps zero_less_mult_iff)
+    finally have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . }
   note * = this
   show ?thesis
     unfolding conditional_entropy_def
     unfolding conditional_mutual_information_generic_eq[OF MX MX MZ]
     by (auto simp: setsum_cartesian_product' setsum_negf[symmetric]
-                   setsum_commute[of _ "space MZ"] *   simp del: divide_pextreal_def
+                   setsum_commute[of _ "space MZ"] *
              intro!: setsum_cong)
 qed
 
@@ -858,29 +849,27 @@
   assumes "simple_function M X" "simple_function M Z"
   shows "\<H>(X | Z) =
      - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
-         real (joint_distribution X Z {(x, z)}) *
-         log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
-  using conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
-  by simp
+         joint_distribution X Z {(x, z)} *
+         log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
+  by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
+     simp
 
 lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
   shows "\<H>(X | Y) =
-    -(\<Sum>y\<in>Y`space M. real (distribution Y {y}) *
-      (\<Sum>x\<in>X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) *
-              log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))"
+    -(\<Sum>y\<in>Y`space M. distribution Y {y} *
+      (\<Sum>x\<in>X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} *
+              log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))"
   unfolding conditional_entropy_eq[OF assms]
-  using finite_distribution_finite[OF finite_random_variable_pairI[OF assms[THEN simple_function_imp_finite_random_variable]]]
   using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]]
-  using finite_distribution_finite[OF Y[THEN simple_function_imp_finite_random_variable]]
-  by (auto simp: setsum_cartesian_product'  setsum_commute[of _ "Y`space M"] setsum_right_distrib real_of_pextreal_eq_0
+  by (auto simp: setsum_cartesian_product'  setsum_commute[of _ "Y`space M"] setsum_right_distrib
            intro!: setsum_cong)
 
 lemma (in information_space) conditional_entropy_eq_cartesian_product:
   assumes "simple_function M X" "simple_function M Y"
   shows "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
-    real (joint_distribution X Y {(x,y)}) *
-    log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))"
+    joint_distribution X Y {(x,y)} *
+    log b (joint_distribution X Y {(x,y)} / distribution Y {y}))"
   unfolding conditional_entropy_eq[OF assms]
   by (auto intro!: setsum_cong simp: setsum_cartesian_product')
 
@@ -890,24 +879,22 @@
   assumes X: "simple_function M X" and Z: "simple_function M Z"
   shows  "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
 proof -
-  let "?XZ x z" = "real (joint_distribution X Z {(x, z)})"
-  let "?Z z" = "real (distribution Z {z})"
-  let "?X x" = "real (distribution X {x})"
+  let "?XZ x z" = "joint_distribution X Z {(x, z)}"
+  let "?Z z" = "distribution Z {z}"
+  let "?X x" = "distribution X {x}"
   note fX = X[THEN simple_function_imp_finite_random_variable]
   note fZ = Z[THEN simple_function_imp_finite_random_variable]
-  note fX[THEN finite_distribution_finite, simp] and fZ[THEN finite_distribution_finite, simp]
   note finite_distribution_order[OF fX fZ, simp]
   { fix x z assume "x \<in> X`space M" "z \<in> Z`space M"
     have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) =
           ?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)"
-      by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff
-                     zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) }
+      by (auto simp: log_simps zero_le_mult_iff field_simps less_le) }
   note * = this
   show ?thesis
     unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z]
-    using setsum_real_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]]
+    using setsum_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]]
     by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric]
-                     setsum_real_distribution)
+                     setsum_distribution)
 qed
 
 lemma (in information_space) conditional_entropy_less_eq_entropy:
@@ -923,21 +910,19 @@
   assumes X: "simple_function M X" and Y: "simple_function M Y"
   shows  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
 proof -
-  let "?XY x y" = "real (joint_distribution X Y {(x, y)})"
-  let "?Y y" = "real (distribution Y {y})"
-  let "?X x" = "real (distribution X {x})"
+  let "?XY x y" = "joint_distribution X Y {(x, y)}"
+  let "?Y y" = "distribution Y {y}"
+  let "?X x" = "distribution X {x}"
   note fX = X[THEN simple_function_imp_finite_random_variable]
   note fY = Y[THEN simple_function_imp_finite_random_variable]
-  note fX[THEN finite_distribution_finite, simp] and fY[THEN finite_distribution_finite, simp]
   note finite_distribution_order[OF fX fY, simp]
   { fix x y assume "x \<in> X`space M" "y \<in> Y`space M"
     have "?XY x y * log b (?XY x y / ?X x) =
           ?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)"
-      by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff
-                     zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) }
+      by (auto simp: log_simps zero_le_mult_iff field_simps less_le) }
   note * = this
   show ?thesis
-    using setsum_real_joint_distribution_singleton[OF fY fX]
+    using setsum_joint_distribution_singleton[OF fY fX]
     unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y]
     unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"]
     by (simp add: * setsum_subtractf setsum_left_distrib[symmetric])
@@ -1063,23 +1048,21 @@
   assumes svi: "subvimage (space M) X P"
   shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
 proof -
-  let "?XP x p" = "real (joint_distribution X P {(x, p)})"
-  let "?X x" = "real (distribution X {x})"
-  let "?P p" = "real (distribution P {p})"
+  let "?XP x p" = "joint_distribution X P {(x, p)}"
+  let "?X x" = "distribution X {x}"
+  let "?P p" = "distribution P {p}"
   note fX = sf(1)[THEN simple_function_imp_finite_random_variable]
   note fP = sf(2)[THEN simple_function_imp_finite_random_variable]
-  note fX[THEN finite_distribution_finite, simp] and fP[THEN finite_distribution_finite, simp]
   note finite_distribution_order[OF fX fP, simp]
-  have "(\<Sum>x\<in>X ` space M. real (distribution X {x}) * log b (real (distribution X {x}))) =
-    (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M.
-    real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})))"
+  have "(\<Sum>x\<in>X ` space M. ?X x * log b (?X x)) =
+    (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M. ?XP x y * log b (?XP x y))"
   proof (subst setsum_image_split[OF svi],
       safe intro!: setsum_mono_zero_cong_left imageI)
     show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)"
       using sf unfolding simple_function_def by auto
   next
     fix p x assume in_space: "p \<in> space M" "x \<in> space M"
-    assume "real (joint_distribution X P {(X x, P p)}) * log b (real (joint_distribution X P {(X x, P p)})) \<noteq> 0"
+    assume "?XP (X x) (P p) * log b (?XP (X x) (P p)) \<noteq> 0"
     hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M \<noteq> {}" by (auto simp: distribution_def)
     with svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`]
     show "x \<in> P -` {P p}" by auto
@@ -1091,20 +1074,16 @@
       by auto
     hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M = X -` {X x} \<inter> space M"
       by auto
-    thus "real (distribution X {X x}) * log b (real (distribution X {X x})) =
-          real (joint_distribution X P {(X x, P p)}) *
-          log b (real (joint_distribution X P {(X x, P p)}))"
+    thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))"
       by (auto simp: distribution_def)
   qed
-  moreover have "\<And>x y. real (joint_distribution X P {(x, y)}) *
-      log b (real (joint_distribution X P {(x, y)}) / real (distribution P {y})) =
-      real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})) -
-      real (joint_distribution X P {(x, y)}) * log b (real (distribution P {y}))"
+  moreover have "\<And>x y. ?XP x y * log b (?XP x y / ?P y) =
+      ?XP x y * log b (?XP x y) - ?XP x y * log b (?P y)"
     by (auto simp add: log_simps zero_less_mult_iff field_simps)
   ultimately show ?thesis
     unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf]
-    using setsum_real_joint_distribution_singleton[OF fX fP]
-    by (simp add: setsum_cartesian_product' setsum_subtractf setsum_real_distribution
+    using setsum_joint_distribution_singleton[OF fX fP]
+    by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution
       setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"])
 qed
 
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -6,6 +6,88 @@
 imports Measure Borel_Space
 begin
 
+lemma extreal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::extreal)"
+  unfolding indicator_def by auto
+
+lemma tendsto_real_max:
+  fixes x y :: real
+  assumes "(X ---> x) net"
+  assumes "(Y ---> y) net"
+  shows "((\<lambda>x. max (X x) (Y x)) ---> max x y) net"
+proof -
+  have *: "\<And>x y :: real. max x y = y + ((x - y) + norm (x - y)) / 2"
+    by (auto split: split_max simp: field_simps)
+  show ?thesis
+    unfolding *
+    by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto
+qed
+
+lemma (in measure_space) measure_Union:
+  assumes "finite S" "S \<subseteq> sets M" "\<And>A B. A \<in> S \<Longrightarrow> B \<in> S \<Longrightarrow> A \<noteq> B \<Longrightarrow> A \<inter> B = {}"
+  shows "setsum \<mu> S = \<mu> (\<Union>S)"
+proof -
+  have "setsum \<mu> S = \<mu> (\<Union>i\<in>S. i)"
+    using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def)
+  also have "\<dots> = \<mu> (\<Union>S)" by (auto intro!: arg_cong[where f=\<mu>])
+  finally show ?thesis .
+qed
+
+lemma (in sigma_algebra) measurable_sets2[intro]:
+  assumes "f \<in> measurable M M'" "g \<in> measurable M M''"
+  and "A \<in> sets M'" "B \<in> sets M''"
+  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
+proof -
+  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+    by auto
+  then show ?thesis using assms by (auto intro: measurable_sets)
+qed
+
+lemma incseq_extreal: "incseq f \<Longrightarrow> incseq (\<lambda>x. extreal (f x))"
+  unfolding incseq_def by auto
+
+lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
+proof
+  assume "\<forall>n. f n \<le> f (Suc n)" then show "incseq f" by (auto intro!: incseq_SucI)
+qed (auto simp: incseq_def)
+
+lemma borel_measurable_real_floor:
+  "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+  unfolding borel.borel_measurable_iff_ge
+proof (intro allI)
+  fix a :: real
+  { fix x have "a \<le> real \<lfloor>x\<rfloor> \<longleftrightarrow> real \<lceil>a\<rceil> \<le> x"
+      using le_floor_eq[of "\<lceil>a\<rceil>" x] ceiling_le_iff[of a "\<lfloor>x\<rfloor>"]
+      unfolding real_eq_of_int by simp }
+  then have "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} = {real \<lceil>a\<rceil>..}" by auto
+  then show "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} \<in> sets borel" by auto
+qed
+
+lemma measure_preservingD2:
+  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
+  unfolding measure_preserving_def by auto
+
+lemma measure_preservingD3:
+  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> space A \<rightarrow> space B"
+  unfolding measure_preserving_def measurable_def by auto
+
+lemma measure_preservingD:
+  "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
+  unfolding measure_preserving_def by auto
+
+lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
+proof -
+  have "\<And>x. real (natfloor (f x)) = max 0 (real \<lfloor>f x\<rfloor>)"
+    by (auto simp: max_def natfloor_def)
+  with borel_measurable_max[OF measurable_comp[OF assms borel_measurable_real_floor] borel_measurable_const]
+  show ?thesis by (simp add: comp_def)
+qed
+
+lemma (in measure_space) AE_not_in:
+  assumes N: "N \<in> null_sets" shows "AE x. x \<notin> N"
+  using N by (rule AE_I') auto
+
 lemma sums_If_finite:
   fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
   assumes finite: "finite {r. P r}"
@@ -55,8 +137,17 @@
     by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def)
 qed
 
+lemma (in sigma_algebra) simple_function_measurable2[intro]:
+  assumes "simple_function M f" "simple_function M g"
+  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
+proof -
+  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+    by auto
+  then show ?thesis using assms[THEN simple_functionD(2)] by auto
+qed
+
 lemma (in sigma_algebra) simple_function_indicator_representation:
-  fixes f ::"'a \<Rightarrow> pextreal"
+  fixes f ::"'a \<Rightarrow> extreal"
   assumes f: "simple_function M f" and x: "x \<in> space M"
   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
   (is "?l = ?r")
@@ -71,7 +162,7 @@
 qed
 
 lemma (in measure_space) simple_function_notspace:
-  "simple_function M (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h")
+  "simple_function M (\<lambda>x. h x * indicator (- space M) x::extreal)" (is "simple_function M ?h")
 proof -
   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
   hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
@@ -111,16 +202,22 @@
 qed
 
 lemma (in sigma_algebra) simple_function_borel_measurable:
-  fixes f :: "'a \<Rightarrow> 'x::t2_space"
+  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
   assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
   shows "simple_function M f"
   using assms unfolding simple_function_def
   by (auto intro: borel_measurable_vimage)
 
+lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+  fixes f :: "'a \<Rightarrow> extreal"
+  shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
+  using simple_function_borel_measurable[of f]
+    borel_measurable_simple_function[of f]
+  by (fastsimp simp: simple_function_def)
+
 lemma (in sigma_algebra) simple_function_const[intro, simp]:
   "simple_function M (\<lambda>x. c)"
   by (auto intro: finite_subset simp: simple_function_def)
-
 lemma (in sigma_algebra) simple_function_compose[intro, simp]:
   assumes "simple_function M f"
   shows "simple_function M (g \<circ> f)"
@@ -189,6 +286,7 @@
   and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
   and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
   and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
+  and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
 
 lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
   assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
@@ -197,247 +295,168 @@
   assume "finite P" from this assms show ?thesis by induct auto
 qed auto
 
-lemma (in sigma_algebra) simple_function_le_measurable:
-  assumes "simple_function M f" "simple_function M g"
-  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
-proof -
-  have *: "{x \<in> space M. f x \<le> g x} =
-    (\<Union>(F, G)\<in>f`space M \<times> g`space M.
-      if F \<le> G then (f -` {F} \<inter> space M) \<inter> (g -` {G} \<inter> space M) else {})"
-    apply (auto split: split_if_asm)
-    apply (rule_tac x=x in bexI)
-    apply (rule_tac x=x in bexI)
-    by simp_all
-  have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow>
-    (f -` {f x} \<inter> space M) \<inter> (g -` {g y} \<inter> space M) \<in> sets M"
-    using assms unfolding simple_function_def by auto
-  have "finite (f`space M \<times> g`space M)"
-    using assms unfolding simple_function_def by auto
-  thus ?thesis unfolding *
-    apply (rule finite_UN)
-    using assms unfolding simple_function_def
-    by (auto intro!: **)
-qed
+lemma (in sigma_algebra)
+  fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
+  shows simple_function_extreal[intro, simp]: "simple_function M (\<lambda>x. extreal (f x))"
+  by (auto intro!: simple_function_compose1[OF sf])
+
+lemma (in sigma_algebra)
+  fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
+  shows simple_function_real_of_nat[intro, simp]: "simple_function M (\<lambda>x. real (f x))"
+  by (auto intro!: simple_function_compose1[OF sf])
 
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
-  fixes u :: "'a \<Rightarrow> pextreal"
+  fixes u :: "'a \<Rightarrow> extreal"
   assumes u: "u \<in> borel_measurable M"
-  shows "\<exists>f. (\<forall>i. simple_function M (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
+  shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
+             (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
 proof -
-  have "\<exists>f. \<forall>x j. (of_nat j \<le> u x \<longrightarrow> f x j = j*2^j) \<and>
-    (u x < of_nat j \<longrightarrow> of_nat (f x j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f x j)))"
-    (is "\<exists>f. \<forall>x j. ?P x j (f x j)")
-  proof(rule choice, rule, rule choice, rule)
-    fix x j show "\<exists>n. ?P x j n"
-    proof cases
-      assume *: "u x < of_nat j"
-      then obtain r where r: "u x = Real r" "0 \<le> r" by (cases "u x") auto
-      from reals_Archimedean6a[of "r * 2^j"]
-      obtain n :: nat where "real n \<le> r * 2 ^ j" "r * 2 ^ j < real (Suc n)"
-        using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff)
-      thus ?thesis using r * by (auto intro!: exI[of _ n])
-    qed auto
-  qed
-  then obtain f where top: "\<And>j x. of_nat j \<le> u x \<Longrightarrow> f x j = j*2^j" and
-    upper: "\<And>j x. u x < of_nat j \<Longrightarrow> of_nat (f x j) \<le> u x * 2^j" and
-    lower: "\<And>j x. u x < of_nat j \<Longrightarrow> u x * 2^j < of_nat (Suc (f x j))" by blast
-
-  { fix j x P
-    assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)"
-    assume 2: "\<And>k. \<lbrakk> u x < of_nat j ; of_nat k \<le> u x * 2^j ; u x * 2^j < of_nat (Suc k) \<rbrakk> \<Longrightarrow> P k"
-    have "P (f x j)"
-    proof cases
-      assume "of_nat j \<le> u x" thus "P (f x j)"
-        using top[of j x] 1 by auto
-    next
-      assume "\<not> of_nat j \<le> u x"
-      hence "u x < of_nat j" "of_nat (f x j) \<le> u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))"
-        using upper lower by auto
-      from 2[OF this] show "P (f x j)" .
-    qed }
-  note fI = this
-
-  { fix j x
-    have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x"
-      by (rule fI, simp, cases "u x") (auto split: split_if_asm) }
-  note f_eq = this
-
-  { fix j x
-    have "f x j \<le> j * 2 ^ j"
-    proof (rule fI)
-      fix k assume *: "u x < of_nat j"
-      assume "of_nat k \<le> u x * 2 ^ j"
-      also have "\<dots> \<le> of_nat (j * 2^j)"
-        using * by (cases "u x") (auto simp: zero_le_mult_iff)
-      finally show "k \<le> j*2^j" by (auto simp del: real_of_nat_mult)
-    qed simp }
+  def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)"
+  { fix x j have "f x j \<le> j * 2 ^ j" unfolding f_def
+    proof (split split_if, intro conjI impI)
+      assume "\<not> real j \<le> u x"
+      then have "natfloor (real (u x) * 2 ^ j) \<le> natfloor (j * 2 ^ j)"
+         by (cases "u x") (auto intro!: natfloor_mono simp: mult_nonneg_nonneg)
+      moreover have "real (natfloor (j * 2 ^ j)) \<le> j * 2^j"
+        by (intro real_natfloor_le) (auto simp: mult_nonneg_nonneg)
+      ultimately show "natfloor (real (u x) * 2 ^ j) \<le> j * 2 ^ j"
+        unfolding real_of_nat_le_iff by auto
+    qed auto }
   note f_upper = this
 
-  let "?g j x" = "of_nat (f x j) / 2^j :: pextreal"
-  show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def
-  proof (safe intro!: exI[of _ ?g])
-    fix j
-    have *: "?g j ` space M \<subseteq> (\<lambda>x. of_nat x / 2^j) ` {..j * 2^j}"
-      using f_upper by auto
-    thus "finite (?g j ` space M)" by (rule finite_subset) auto
-  next
-    fix j t assume "t \<in> space M"
-    have **: "?g j -` {?g j t} \<inter> space M = {x \<in> space M. f t j = f x j}"
-      by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff)
+  have real_f:
+    "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))"
+    unfolding f_def by auto
 
-    show "?g j -` {?g j t} \<inter> space M \<in> sets M"
-    proof cases
-      assume "of_nat j \<le> u t"
-      hence "?g j -` {?g j t} \<inter> space M = {x\<in>space M. of_nat j \<le> u x}"
-        unfolding ** f_eq[symmetric] by auto
-      thus "?g j -` {?g j t} \<inter> space M \<in> sets M"
-        using u by auto
-    next
-      assume not_t: "\<not> of_nat j \<le> u t"
-      hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \<le> j*2^j` by auto
-      have split_vimage: "?g j -` {?g j t} \<inter> space M =
-          {x\<in>space M. of_nat (f t j)/2^j \<le> u x} \<inter> {x\<in>space M. u x < of_nat (Suc (f t j))/2^j}"
-        unfolding **
-      proof safe
-        fix x assume [simp]: "f t j = f x j"
-        have *: "\<not> of_nat j \<le> u x" using not_t f_eq[symmetric] by simp
-        hence "of_nat (f t j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f t j))"
-          using upper lower by auto
-        hence "of_nat (f t j) / 2^j \<le> u x \<and> u x < of_nat (Suc (f t j))/2^j" using *
-          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
-        thus "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" by auto
+  let "?g j x" = "real (f x j) / 2^j :: extreal"
+  show ?thesis
+  proof (intro exI[of _ ?g] conjI allI ballI)
+    fix i
+    have "simple_function M (\<lambda>x. real (f x i))"
+    proof (intro simple_function_borel_measurable)
+      show "(\<lambda>x. real (f x i)) \<in> borel_measurable M"
+        using u by (auto intro!: measurable_If simp: real_f)
+      have "(\<lambda>x. real (f x i))`space M \<subseteq> real`{..i*2^i}"
+        using f_upper[of _ i] by auto
+      then show "finite ((\<lambda>x. real (f x i))`space M)"
+        by (rule finite_subset) auto
+    qed
+    then show "simple_function M (?g i)"
+      by (auto intro: simple_function_extreal simple_function_div)
+  next
+    show "incseq ?g"
+    proof (intro incseq_extreal incseq_SucI le_funI)
+      fix x and i :: nat
+      have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
+      proof ((split split_if)+, intro conjI impI)
+        assume "extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+        then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)"
+          by (cases "u x") (auto intro!: le_natfloor)
       next
-        fix x
-        assume "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j"
-        hence "of_nat (f t j) \<le> u x * 2 ^ j \<and> u x * 2 ^ j < of_nat (Suc (f t j))"
-          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
-        hence 1: "of_nat (f t j) \<le> u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto
-        note 2
-        also have "\<dots> \<le> of_nat (j*2^j)"
-          using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult)
-        finally have bound_ux: "u x < of_nat j"
-          by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq)
-        show "f t j = f x j"
-        proof (rule antisym)
-          from 1 lower[OF bound_ux]
-          show "f t j \<le> f x j" by (cases "u x") (auto split: split_if_asm)
-          from upper[OF bound_ux] 2
-          show "f x j \<le> f t j" by (cases "u x") (auto split: split_if_asm)
+        assume "\<not> extreal (real i) \<le> u x" "extreal (real (Suc i)) \<le> u x"
+        then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i"
+          by (cases "u x") auto
+      next
+        assume "\<not> extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+        have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2"
+          by simp
+        also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)"
+        proof cases
+          assume "0 \<le> u x" then show ?thesis
+            by (intro le_mult_natfloor) (cases "u x", auto intro!: mult_nonneg_nonneg)
+        next
+          assume "\<not> 0 \<le> u x" then show ?thesis
+            by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg)
         qed
-      qed
-      show ?thesis unfolding split_vimage using u by auto
+        also have "\<dots> = natfloor (real (u x) * 2 ^ Suc i)"
+          by (simp add: ac_simps)
+        finally show "natfloor (real (u x) * 2 ^ i) * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)" .
+      qed simp
+      then show "?g i x \<le> ?g (Suc i) x"
+        by (auto simp: field_simps)
     qed
   next
-    fix j t assume "?g j t = \<omega>" thus False by (auto simp: power_le_zero_eq)
-  next
-    fix t
-    { fix i
-      have "f t i * 2 \<le> f t (Suc i)"
-      proof (rule fI)
-        assume "of_nat (Suc i) \<le> u t"
-        hence "of_nat i \<le> u t" by (cases "u t") auto
-        thus "f t i * 2 \<le> Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp
-      next
-        fix k
-        assume *: "u t * 2 ^ Suc i < of_nat (Suc k)"
-        show "f t i * 2 \<le> k"
-        proof (rule fI)
-          assume "of_nat i \<le> u t"
-          hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i"
-            by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
-          also have "\<dots> < of_nat (Suc k)" using * by auto
-          finally show "i * 2 ^ i * 2 \<le> k"
-            by (auto simp del: real_of_nat_mult)
-        next
-          fix j assume "of_nat j \<le> u t * 2 ^ i"
-          with * show "j * 2 \<le> k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
+    fix x show "(SUP i. ?g i x) = max 0 (u x)"
+    proof (rule extreal_SUPI)
+      fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
+        by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
+                                     mult_nonpos_nonneg mult_nonneg_nonneg)
+    next
+      fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> ?g i x \<le> y"
+      have "\<And>i. 0 \<le> ?g i x" by (auto simp: divide_nonneg_pos)
+      from order_trans[OF this *] have "0 \<le> y" by simp
+      show "max 0 (u x) \<le> y"
+      proof (cases y)
+        case (real r)
+        with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
+        from real_arch_lt[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
+        then have "\<exists>p. max 0 (u x) = extreal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
+        then guess p .. note ux = this
+        obtain m :: nat where m: "p < real m" using real_arch_lt ..
+        have "p \<le> r"
+        proof (rule ccontr)
+          assume "\<not> p \<le> r"
+          with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"]
+          obtain N where "\<forall>n\<ge>N. r * 2^n < p * 2^n - 1" by (auto simp: inverse_eq_divide field_simps)
+          then have "r * 2^max N m < p * 2^max N m - 1" by simp
+          moreover
+          have "real (natfloor (p * 2 ^ max N m)) \<le> r * 2 ^ max N m"
+            using *[of "max N m"] m unfolding real_f using ux
+            by (cases "0 \<le> u x") (simp_all add: max_def mult_nonneg_nonneg split: split_if_asm)
+          then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m"
+            by (metis real_natfloor_gt_diff_one less_le_trans)
+          ultimately show False by auto
         qed
-      qed
-      thus "?g i t \<le> ?g (Suc i) t"
-        by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) }
-    hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto
+        then show "max 0 (u x) \<le> y" using real ux by simp
+      qed (insert `0 \<le> y`, auto)
+    qed
+  qed (auto simp: divide_nonneg_pos)
+qed
 
-    show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
-    proof (rule pextreal_SUPI)
-      fix j show "of_nat (f t j) / 2 ^ j \<le> u t"
-      proof (rule fI)
-        assume "of_nat j \<le> u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \<le> u t"
-          by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps)
-      next
-        fix k assume "of_nat k \<le> u t * 2 ^ j"
-        thus "of_nat k / 2 ^ j \<le> u t"
-          by (cases "u t")
-             (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff)
-      qed
-    next
-      fix y :: pextreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
-      show "u t \<le> y"
-      proof (cases "u t")
-        case (preal r)
-        show ?thesis
-        proof (rule ccontr)
-          assume "\<not> u t \<le> y"
-          then obtain p where p: "y = Real p" "0 \<le> p" "p < r" using preal by (cases y) auto
-          with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"]
-          obtain n where n: "\<And>N. n \<le> N \<Longrightarrow> inverse (2^N) < r - p" by auto
-          let ?N = "max n (natfloor r + 1)"
-          have "u t < of_nat ?N" "n \<le> ?N"
-            using ge_natfloor_plus_one_imp_gt[of r n] preal
-            using real_natfloor_add_one_gt
-            by (auto simp: max_def real_of_nat_Suc)
-          from lower[OF this(1)]
-          have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq
-            using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm)
-          hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N"
-            using preal by (auto simp: field_simps divide_real_def[symmetric])
-          with n[OF `n \<le> ?N`] p preal *[of ?N]
-          show False
-            by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm)
-        qed
-      next
-        case infinite
-        { fix j have "f t j = j*2^j" using top[of j t] infinite by simp
-          hence "of_nat j \<le> y" using *[of j]
-            by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) }
-        note all_less_y = this
-        show ?thesis unfolding infinite
-        proof (rule ccontr)
-          assume "\<not> \<omega> \<le> y" then obtain r where r: "y = Real r" "0 \<le> r" by (cases y) auto
-          moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat)
-          with all_less_y[of n] r show False by auto
-        qed
-      qed
-    qed
+lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
+  fixes u :: "'a \<Rightarrow> extreal"
+  assumes u: "u \<in> borel_measurable M"
+  obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
+    "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
+  using borel_measurable_implies_simple_function_sequence[OF u] by auto
+
+lemma (in sigma_algebra) simple_function_If_set:
+  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
+  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
+proof -
+  def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
+  show ?thesis unfolding simple_function_def
+  proof safe
+    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
+    from finite_subset[OF this] assms
+    show "finite (?IF ` space M)" unfolding simple_function_def by auto
+  next
+    fix x assume "x \<in> space M"
+    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
+      then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
+      else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
+      using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
+    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
+      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
+    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
   qed
 qed
 
-lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
-  fixes u :: "'a \<Rightarrow> pextreal"
-  assumes "u \<in> borel_measurable M"
-  obtains (x) f where "f \<up> u" "\<And>i. simple_function M (f i)" "\<And>i. \<omega>\<notin>f i`space M"
+lemma (in sigma_algebra) simple_function_If:
+  assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
+  shows "simple_function M (\<lambda>x. if P x then f x else g x)"
 proof -
-  from borel_measurable_implies_simple_function_sequence[OF assms]
-  obtain f where x: "\<And>i. simple_function M (f i)" "f \<up> u"
-    and fin: "\<And>i. \<And>x. x\<in>space M \<Longrightarrow> f i x \<noteq> \<omega>" by auto
-  { fix i from fin[of _ i] have "\<omega> \<notin> f i`space M" by fastsimp }
-  with x show thesis by (auto intro!: that[of f])
+  have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
+  with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
 qed
 
-lemma (in sigma_algebra) simple_function_eq_borel_measurable:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  shows "simple_function M f \<longleftrightarrow>
-    finite (f`space M) \<and> f \<in> borel_measurable M"
-  using simple_function_borel_measurable[of f]
-    borel_measurable_simple_function[of f]
-  by (fastsimp simp: simple_function_def)
-
 lemma (in measure_space) simple_function_restricted:
-  fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M"
+  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
   shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
     (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
 proof -
   interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
-  have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
+  have f: "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
   proof cases
     assume "A = space M"
     then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
@@ -456,7 +475,7 @@
         using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
     next
       fix x
-      assume "indicator A x \<noteq> (0::pextreal)"
+      assume "indicator A x \<noteq> (0::extreal)"
       then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
       moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
       ultimately show "f x = 0" by auto
@@ -467,7 +486,8 @@
     unfolding simple_function_eq_borel_measurable
       R.simple_function_eq_borel_measurable
     unfolding borel_measurable_restricted[OF `A \<in> sets M`]
-    by auto
+    using assms(1)[THEN sets_into_space]
+    by (auto simp: indicator_def)
 qed
 
 lemma (in sigma_algebra) simple_function_subalgebra:
@@ -504,7 +524,7 @@
   "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
 
 syntax
-  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
+  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
 
 translations
   "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
@@ -540,7 +560,7 @@
 qed
 
 lemma (in measure_space) simple_function_partition:
-  assumes "simple_function M f" and "simple_function M g"
+  assumes f: "simple_function M f" and g: "simple_function M g"
   shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)"
     (is "_ = setsum _ (?p ` space M)")
 proof-
@@ -559,23 +579,16 @@
     hence "finite (?p ` (A \<inter> space M))"
       by (rule finite_subset) auto }
   note this[intro, simp]
+  note sets = simple_function_measurable2[OF f g]
 
   { fix x assume "x \<in> space M"
     have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
-    moreover {
-      fix x y
-      have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M
-          = (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" by auto
-      assume "x \<in> space M" "y \<in> space M"
-      hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M"
-        using assms unfolding simple_function_def * by auto }
-    ultimately
-    have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
-      by (subst measure_finitely_additive) auto }
+    with sets have "\<mu> (f -` {f x} \<inter> space M) = setsum \<mu> (?sub (f x))"
+      by (subst measure_Union) auto }
   hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
-    unfolding simple_integral_def
-    by (subst setsum_Sigma[symmetric],
-       auto intro!: setsum_cong simp: setsum_right_distrib[symmetric])
+    unfolding simple_integral_def using f sets
+    by (subst setsum_Sigma[symmetric])
+       (auto intro!: setsum_cong setsum_extreal_right_distrib)
   also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
   proof -
     have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
@@ -595,7 +608,7 @@
 qed
 
 lemma (in measure_space) simple_integral_add[simp]:
-  assumes "simple_function M f" and "simple_function M g"
+  assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
   shows "(\<integral>\<^isup>Sx. f x + g x \<partial>M) = integral\<^isup>S M f + integral\<^isup>S M g"
 proof -
   { fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
@@ -603,63 +616,43 @@
     hence "(\<lambda>a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}"
         "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
       by auto }
-  thus ?thesis
+  with assms show ?thesis
     unfolding
-      simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
-      simple_function_partition[OF `simple_function M f` `simple_function M g`]
-      simple_function_partition[OF `simple_function M g` `simple_function M f`]
-    apply (subst (3) Int_commute)
-    by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong)
+      simple_function_partition[OF simple_function_add[OF f g] simple_function_Pair[OF f g]]
+      simple_function_partition[OF f g]
+      simple_function_partition[OF g f]
+    by (subst (3) Int_commute)
+       (auto simp add: extreal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
 qed
 
 lemma (in measure_space) simple_integral_setsum[simp]:
+  assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
   assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
   shows "(\<integral>\<^isup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>S M (f i))"
 proof cases
   assume "finite P"
   from this assms show ?thesis
-    by induct (auto simp: simple_function_setsum simple_integral_add)
+    by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
 qed auto
 
 lemma (in measure_space) simple_integral_mult[simp]:
-  assumes "simple_function M f"
+  assumes f: "simple_function M f" "\<And>x. 0 \<le> f x"
   shows "(\<integral>\<^isup>Sx. c * f x \<partial>M) = c * integral\<^isup>S M f"
 proof -
-  note mult = simple_function_mult[OF simple_function_const[of c] assms]
+  note mult = simple_function_mult[OF simple_function_const[of c] f(1)]
   { fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
     assume "x \<in> space M"
     hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
       by auto }
-  thus ?thesis
-    unfolding simple_function_partition[OF mult assms]
-      simple_function_partition[OF assms mult]
-    by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
-qed
-
-lemma (in sigma_algebra) simple_function_If:
-  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<in> sets M"
-  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
-proof -
-  def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
-  show ?thesis unfolding simple_function_def
-  proof safe
-    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
-    from finite_subset[OF this] assms
-    show "finite (?IF ` space M)" unfolding simple_function_def by auto
-  next
-    fix x assume "x \<in> space M"
-    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
-      then ((F (f x) \<inter> A) \<union> (G (f x) - (G (f x) \<inter> A)))
-      else ((F (g x) \<inter> A) \<union> (G (g x) - (G (g x) \<inter> A))))"
-      using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
-    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
-      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
-    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
-  qed
+  with assms show ?thesis
+    unfolding simple_function_partition[OF mult f(1)]
+              simple_function_partition[OF f(1) mult]
+    by (subst setsum_extreal_right_distrib)
+       (auto intro!: extreal_0_le_mult setsum_cong simp: mult_assoc)
 qed
 
 lemma (in measure_space) simple_integral_mono_AE:
-  assumes "simple_function M f" and "simple_function M g"
+  assumes f: "simple_function M f" and g: "simple_function M g"
   and mono: "AE x. f x \<le> g x"
   shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
 proof -
@@ -668,14 +661,16 @@
     "\<And>x. f -` {f x} \<inter> g -` {g x} \<inter> space M = ?S x" by auto
   show ?thesis
     unfolding *
-      simple_function_partition[OF `simple_function M f` `simple_function M g`]
-      simple_function_partition[OF `simple_function M g` `simple_function M f`]
+      simple_function_partition[OF f g]
+      simple_function_partition[OF g f]
   proof (safe intro!: setsum_mono)
     fix x assume "x \<in> space M"
     then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto
     show "the_elem (f`?S x) * \<mu> (?S x) \<le> the_elem (g`?S x) * \<mu> (?S x)"
     proof (cases "f x \<le> g x")
-      case True then show ?thesis using * by (auto intro!: mult_right_mono)
+      case True then show ?thesis
+        using * assms(1,2)[THEN simple_functionD(2)]
+        by (auto intro!: extreal_mult_right_mono)
     next
       case False
       obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
@@ -685,7 +680,10 @@
         by (rule_tac Int) (auto intro!: simple_functionD)
       ultimately have "\<mu> (?S x) \<le> \<mu> N"
         using `N \<in> sets M` by (auto intro!: measure_mono)
-      then show ?thesis using `\<mu> N = 0` by auto
+      moreover have "0 \<le> \<mu> (?S x)"
+        using assms(1,2)[THEN simple_functionD(2)] by auto
+      ultimately have "\<mu> (?S x) = 0" using `\<mu> N = 0` by auto
+      then show ?thesis by simp
     qed
   qed
 qed
@@ -697,7 +695,8 @@
   using assms by (intro simple_integral_mono_AE) auto
 
 lemma (in measure_space) simple_integral_cong_AE:
-  assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x"
+  assumes "simple_function M f" and "simple_function M g"
+  and "AE x. f x = g x"
   shows "integral\<^isup>S M f = integral\<^isup>S M g"
   using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
 
@@ -765,7 +764,7 @@
   assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
   thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
 next
-  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pextreal}" by auto
+  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::extreal}" by auto
   thus ?thesis
     using simple_integral_indicator[OF assms simple_function_const[of 1]]
     using sets_into_space[OF assms]
@@ -773,13 +772,13 @@
 qed
 
 lemma (in measure_space) simple_integral_null_set:
-  assumes "simple_function M u" "N \<in> null_sets"
+  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
   shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
 proof -
-  have "AE x. indicator N x = (0 :: pextreal)"
+  have "AE x. indicator N x = (0 :: extreal)"
     using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
   then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
-    using assms by (intro simple_integral_cong_AE) auto
+    using assms apply (intro simple_integral_cong_AE) by auto
   then show ?thesis by simp
 qed
 
@@ -813,7 +812,7 @@
     by (auto simp: indicator_def split: split_if_asm)
   then show "f x * \<mu> (f -` {f x} \<inter> A) =
     f x * \<mu> (?f -` {f x} \<inter> space M)"
-    unfolding pextreal_mult_cancel_left by auto
+    unfolding extreal_mult_cancel_left by auto
 qed
 
 lemma (in measure_space) simple_integral_subalgebra:
@@ -821,10 +820,6 @@
   shows "integral\<^isup>S N = integral\<^isup>S M"
   unfolding simple_integral_def_raw by simp
 
-lemma measure_preservingD:
-  "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
-  unfolding measure_preserving_def by auto
-
 lemma (in measure_space) simple_integral_vimage:
   assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
     and f: "simple_function M' f"
@@ -853,196 +848,164 @@
   qed
 qed
 
+lemma (in measure_space) simple_integral_cmult_indicator:
+  assumes A: "A \<in> sets M"
+  shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * \<mu> A"
+  using simple_integral_mult[OF simple_function_indicator[OF A]]
+  unfolding simple_integral_indicator_only[OF A] by simp
+
+lemma (in measure_space) simple_integral_positive:
+  assumes f: "simple_function M f" and ae: "AE x. 0 \<le> f x"
+  shows "0 \<le> integral\<^isup>S M f"
+proof -
+  have "integral\<^isup>S M (\<lambda>x. 0) \<le> integral\<^isup>S M f"
+    using simple_integral_mono_AE[OF _ f ae] by auto
+  then show ?thesis by simp
+qed
+
 section "Continuous positive integration"
 
 definition positive_integral_def:
-  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^isup>S M g)"
+  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
 
 syntax
-  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> pextreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> pextreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
+  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
 
 translations
   "\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
 
-lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f =
-    (SUP g : {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. integral\<^isup>S M g)"
-  (is "_ = ?alt")
-proof (rule antisym SUP_leI)
-  show "integral\<^isup>P M f \<le> ?alt" unfolding positive_integral_def
-  proof (safe intro!: SUP_leI)
-    fix g assume g: "simple_function M g" "g \<le> f"
-    let ?G = "g -` {\<omega>} \<inter> space M"
-    show "integral\<^isup>S M g \<le>
-      (SUP h : {i. simple_function M i \<and> i \<le> f \<and> \<omega> \<notin> i ` space M}. integral\<^isup>S M h)"
-      (is "integral\<^isup>S M g \<le> SUPR ?A _")
-    proof cases
-      let ?g = "\<lambda>x. indicator (space M - ?G) x * g x"
-      have g': "simple_function M ?g"
-        using g by (auto intro: simple_functionD)
-      moreover
-      assume "\<mu> ?G = 0"
-      then have "AE x. g x = ?g x" using g
-        by (intro AE_I[where N="?G"])
-           (auto intro: simple_functionD simp: indicator_def)
-      with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g"
-        by (rule simple_integral_cong_AE)
-      moreover have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
-      from this `g \<le> f` have "?g \<le> f" by (rule order_trans)
-      moreover have "\<omega> \<notin> ?g ` space M"
-        by (auto simp: indicator_def split: split_if_asm)
-      ultimately show ?thesis by (auto intro!: le_SUPI)
-    next
-      assume "\<mu> ?G \<noteq> 0"
-      then have "?G \<noteq> {}" by auto
-      then have "\<omega> \<in> g`space M" by force
-      then have "space M \<noteq> {}" by auto
-      have "SUPR ?A (integral\<^isup>S M) = \<omega>"
-      proof (intro SUP_\<omega>[THEN iffD2] allI impI)
-        fix x assume "x < \<omega>"
-        then guess n unfolding less_\<omega>_Ex_of_nat .. note n = this
-        then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp
-        let ?g = "\<lambda>x. (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * indicator ?G x"
-        show "\<exists>i\<in>?A. x < integral\<^isup>S M i"
-        proof (intro bexI impI CollectI conjI)
-          show "simple_function M ?g" using g
-            by (auto intro!: simple_functionD simple_function_add)
-          have "?g \<le> g" by (auto simp: le_fun_def indicator_def)
-          from this g(2) show "?g \<le> f" by (rule order_trans)
-          show "\<omega> \<notin> ?g ` space M"
-            using `\<mu> ?G \<noteq> 0` by (auto simp: indicator_def split: split_if_asm)
-          have "x < (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * \<mu> ?G"
-            using n `\<mu> ?G \<noteq> 0` `0 < n`
-            by (auto simp: pextreal_noteq_omega_Ex field_simps)
-          also have "\<dots> = integral\<^isup>S M ?g" using g `space M \<noteq> {}`
-            by (subst simple_integral_indicator)
-               (auto simp: image_constant ac_simps dest: simple_functionD)
-          finally show "x < integral\<^isup>S M ?g" .
-        qed
-      qed
-      then show ?thesis by simp
-    qed
-  qed
-qed (auto intro!: SUP_subset simp: positive_integral_def)
-
 lemma (in measure_space) positive_integral_cong_measure:
   assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
   shows "integral\<^isup>P N f = integral\<^isup>P M f"
-proof -
-  interpret v: measure_space N
-    by (rule measure_space_cong) fact+
-  with assms show ?thesis
-    unfolding positive_integral_def SUPR_def
-    by (auto intro!: arg_cong[where f=Sup] image_cong
-             simp: simple_integral_cong_measure[OF assms]
-                   simple_function_cong_algebra[OF assms(2,3)])
-qed
+  unfolding positive_integral_def
+  unfolding simple_function_cong_algebra[OF assms(2,3), symmetric]
+  using AE_cong_measure[OF assms]
+  using simple_integral_cong_measure[OF assms]
+  by (auto intro!: SUP_cong)
+
+lemma (in measure_space) positive_integral_positive:
+  "0 \<le> integral\<^isup>P M f"
+  by (auto intro!: le_SUPI2[of "\<lambda>x. 0"] simp: positive_integral_def le_fun_def)
 
-lemma (in measure_space) positive_integral_alt1:
-  "integral\<^isup>P M f =
-    (SUP g : {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. integral\<^isup>S M g)"
-  unfolding positive_integral_alt SUPR_def
-proof (safe intro!: arg_cong[where f=Sup])
-  fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x"
-  assume "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
-  hence "?g \<le> f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g"
-    "\<omega> \<notin> g`space M"
-    unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong)
-  thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
-    by auto
-next
-  fix g assume "simple_function M g" "g \<le> f" "\<omega> \<notin> g`space M"
-  hence "simple_function M g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
-    by (auto simp add: le_fun_def image_iff)
-  thus "integral\<^isup>S M g \<in> integral\<^isup>S M ` {g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
-    by auto
-qed
+lemma (in measure_space) positive_integral_def_finite:
+  "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^isup>S M g)"
+    (is "_ = SUPR ?A ?f")
+  unfolding positive_integral_def
+proof (safe intro!: antisym SUP_leI)
+  fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f"
+  let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
+  note gM = g(1)[THEN borel_measurable_simple_function]
+  have \<mu>G_pos: "0 \<le> \<mu> ?G" using gM by auto
+  let "?g y x" = "if g x = \<infinity> then y else max 0 (g x)"
+  from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A"
+    apply (safe intro!: simple_function_max simple_function_If)
+    apply (force simp: max_def le_fun_def split: split_if_asm)+
+    done
+  show "integral\<^isup>S M g \<le> SUPR ?A ?f"
+  proof cases
+    have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
+    assume "\<mu> ?G = 0"
+    with gM have "AE x. x \<notin> ?G" by (simp add: AE_iff_null_set)
+    with gM g show ?thesis
+      by (intro le_SUPI2[OF g0] simple_integral_mono_AE)
+         (auto simp: max_def intro!: simple_function_If)
+  next
+    assume \<mu>G: "\<mu> ?G \<noteq> 0"
+    have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
+    proof (intro SUP_PInfty)
+      fix n :: nat
+      let ?y = "extreal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
+      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: extreal_divide_eq)
+      then have "?g ?y \<in> ?A" by (rule g_in_A)
+      have "real n \<le> ?y * \<mu> ?G"
+        using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
+      also have "\<dots> = (\<integral>\<^isup>Sx. ?y * indicator ?G x \<partial>M)"
+        using `0 \<le> ?y` `?g ?y \<in> ?A` gM
+        by (subst simple_integral_cmult_indicator) auto
+      also have "\<dots> \<le> integral\<^isup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
+        by (intro simple_integral_mono) auto
+      finally show "\<exists>i\<in>?A. real n \<le> integral\<^isup>S M i"
+        using `?g ?y \<in> ?A` by blast
+    qed
+    then show ?thesis by simp
+  qed
+qed (auto intro: le_SUPI)
 
-lemma (in measure_space) positive_integral_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
-  shows "integral\<^isup>P M f = integral\<^isup>P M g"
-proof -
-  have "\<And>h. (\<forall>x\<in>space M. h x \<le> f x \<and> h x \<noteq> \<omega>) = (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>)"
-    using assms by auto
-  thus ?thesis unfolding positive_integral_alt1 by auto
+lemma (in measure_space) positive_integral_mono_AE:
+  assumes ae: "AE x. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
+  unfolding positive_integral_def
+proof (safe intro!: SUP_mono)
+  fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
+  from ae[THEN AE_E] guess N . note N = this
+  then have ae_N: "AE x. x \<notin> N" by (auto intro: AE_not_in)
+  let "?n x" = "n x * indicator (space M - N) x"
+  have "AE x. n x \<le> ?n x" "simple_function M ?n"
+    using n N ae_N by auto
+  moreover
+  { fix x have "?n x \<le> max 0 (v x)"
+    proof cases
+      assume x: "x \<in> space M - N"
+      with N have "u x \<le> v x" by auto
+      with n(2)[THEN le_funD, of x] x show ?thesis
+        by (auto simp: max_def split: split_if_asm)
+    qed simp }
+  then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI)
+  moreover have "integral\<^isup>S M n \<le> integral\<^isup>S M ?n"
+    using ae_N N n by (auto intro!: simple_integral_mono_AE)
+  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^isup>S M n \<le> integral\<^isup>S M m"
+    by force
 qed
 
-lemma (in measure_space) positive_integral_eq_simple_integral:
-  assumes "simple_function M f"
-  shows "integral\<^isup>P M f = integral\<^isup>S M f"
-  unfolding positive_integral_def
-proof (safe intro!: pextreal_SUPI)
-  fix g assume "simple_function M g" "g \<le> f"
-  with assms show "integral\<^isup>S M g \<le> integral\<^isup>S M f"
-    by (auto intro!: simple_integral_mono simp: le_fun_def)
-next
-  fix y assume "\<forall>x. x\<in>{g. simple_function M g \<and> g \<le> f} \<longrightarrow> integral\<^isup>S M x \<le> y"
-  with assms show "integral\<^isup>S M f \<le> y" by auto
-qed
-
-lemma (in measure_space) positive_integral_mono_AE:
-  assumes ae: "AE x. u x \<le> v x"
-  shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
-  unfolding positive_integral_alt1
-proof (safe intro!: SUPR_mono)
-  fix a assume a: "simple_function M a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
-  from ae obtain N where N: "{x\<in>space M. \<not> u x \<le> v x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
-    by (auto elim!: AE_E)
-  have "simple_function M (\<lambda>x. a x * indicator (space M - N) x)"
-    using `N \<in> sets M` a by auto
-  with a show "\<exists>b\<in>{g. simple_function M g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}.
-    integral\<^isup>S M a \<le> integral\<^isup>S M b"
-  proof (safe intro!: bexI[of _ "\<lambda>x. a x * indicator (space M - N) x"]
-                      simple_integral_mono_AE)
-    show "AE x. a x \<le> a x * indicator (space M - N) x"
-    proof (rule AE_I, rule subset_refl)
-      have *: "{x \<in> space M. \<not> a x \<le> a x * indicator (space M - N) x} =
-        N \<inter> {x \<in> space M. a x \<noteq> 0}" (is "?N = _")
-        using `N \<in> sets M`[THEN sets_into_space] by (auto simp: indicator_def)
-      then show "?N \<in> sets M"
-        using `N \<in> sets M` `simple_function M a`[THEN borel_measurable_simple_function]
-        by (auto intro!: measure_mono Int)
-      then have "\<mu> ?N \<le> \<mu> N"
-        unfolding * using `N \<in> sets M` by (auto intro!: measure_mono)
-      then show "\<mu> ?N = 0" using `\<mu> N = 0` by auto
-    qed
-  next
-    fix x assume "x \<in> space M"
-    show "a x * indicator (space M - N) x \<le> v x"
-    proof (cases "x \<in> N")
-      case True then show ?thesis by simp
-    next
-      case False
-      with N mono have "a x \<le> u x" "u x \<le> v x" using `x \<in> space M` by auto
-      with False `x \<in> space M` show "a x * indicator (space M - N) x \<le> v x" by auto
-    qed
-    assume "a x * indicator (space M - N) x = \<omega>"
-    with mono `x \<in> space M` show False
-      by (simp split: split_if_asm add: indicator_def)
-  qed
-qed
+lemma (in measure_space) positive_integral_mono:
+  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
+  by (auto intro: positive_integral_mono_AE)
 
 lemma (in measure_space) positive_integral_cong_AE:
   "AE x. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
   by (auto simp: eq_iff intro!: positive_integral_mono_AE)
 
-lemma (in measure_space) positive_integral_mono:
-  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
-  by (auto intro: positive_integral_mono_AE)
+lemma (in measure_space) positive_integral_cong:
+  "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
+  by (auto intro: positive_integral_cong_AE)
 
-lemma image_set_cong:
-  assumes A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. f x = g y"
-  assumes B: "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. g y = f x"
-  shows "f ` A = g ` B"
-  using assms by blast
+lemma (in measure_space) positive_integral_eq_simple_integral:
+  assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+proof -
+  let "?f x" = "f x * indicator (space M) x"
+  have f': "simple_function M ?f" using f by auto
+  with f(2) have [simp]: "max 0 \<circ> ?f = ?f"
+    by (auto simp: fun_eq_iff max_def split: split_indicator)
+  have "integral\<^isup>P M ?f \<le> integral\<^isup>S M ?f" using f'
+    by (force intro!: SUP_leI simple_integral_mono simp: le_fun_def positive_integral_def)
+  moreover have "integral\<^isup>S M ?f \<le> integral\<^isup>P M ?f"
+    unfolding positive_integral_def
+    using f' by (auto intro!: le_SUPI)
+  ultimately show ?thesis
+    by (simp cong: positive_integral_cong simple_integral_cong)
+qed
+
+lemma (in measure_space) positive_integral_eq_simple_integral_AE:
+  assumes f: "simple_function M f" "AE x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+proof -
+  have "AE x. f x = max 0 (f x)" using f by (auto split: split_max)
+  with f have "integral\<^isup>P M f = integral\<^isup>S M (\<lambda>x. max 0 (f x))"
+    by (simp cong: positive_integral_cong_AE simple_integral_cong_AE
+             add: positive_integral_eq_simple_integral)
+  with assms show ?thesis
+    by (auto intro!: simple_integral_cong_AE split: split_max)
+qed
 
 lemma (in measure_space) positive_integral_SUP_approx:
-  assumes "f \<up> s"
-  and f: "\<And>i. f i \<in> borel_measurable M"
-  and "simple_function M u"
-  and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
+  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+  and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
   shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
-proof (rule pextreal_le_mult_one_interval)
-  fix a :: pextreal assume "0 < a" "a < 1"
+proof (rule extreal_le_mult_one_interval)
+  have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
+    using f(3) by (auto intro!: le_SUPI2 positive_integral_positive)
+  then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
+  have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
+    using u(3) by auto
+  fix a :: extreal assume "0 < a" "a < 1"
   hence "a \<noteq> 0" by auto
   let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
   have B: "\<And>i. ?B i \<in> sets M"
@@ -1054,203 +1017,269 @@
     proof safe
       fix i x assume "a * u x \<le> f i x"
       also have "\<dots> \<le> f (Suc i) x"
-        using `f \<up> s` unfolding isoton_def le_fun_def by auto
+        using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto
       finally show "a * u x \<le> f (Suc i) x" .
     qed }
   note B_mono = this
 
-  have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
-    using `simple_function M u` by (auto simp add: simple_function_def)
+  note B_u = Int[OF u(1)[THEN simple_functionD(2)] B]
 
-  have "\<And>i. (\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
-  proof safe
-    fix x i assume "x \<in> space M"
-    show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
-    proof cases
-      assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
-    next
-      assume "u x \<noteq> 0"
-      with `a < 1` real `x \<in> space M`
-      have "a * u x < 1 * u x" by (rule_tac pextreal_mult_strict_right_mono) (auto simp: image_iff)
-      also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
-        unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
-      finally obtain i where "a * u x < f i x" unfolding SUPR_def
-        by (auto simp add: less_Sup_iff)
-      hence "a * u x \<le> f i x" by auto
-      thus ?thesis using `x \<in> space M` by auto
+  let "?B' i n" = "(u -` {i} \<inter> space M) \<inter> ?B n"
+  have measure_conv: "\<And>i. \<mu> (u -` {i} \<inter> space M) = (SUP n. \<mu> (?B' i n))"
+  proof -
+    fix i
+    have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto
+    have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI)
+    have "(\<Union>n. ?B' i n) = u -` {i} \<inter> space M"
+    proof safe
+      fix x i assume x: "x \<in> space M"
+      show "x \<in> (\<Union>i. ?B' (u x) i)"
+      proof cases
+        assume "u x = 0" thus ?thesis using `x \<in> space M` f(3) by simp
+      next
+        assume "u x \<noteq> 0"
+        with `a < 1` u_range[OF `x \<in> space M`]
+        have "a * u x < 1 * u x"
+          by (intro extreal_mult_strict_right_mono) (auto simp: image_iff)
+        also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUPR_apply)
+        finally obtain i where "a * u x < f i x" unfolding SUPR_def
+          by (auto simp add: less_Sup_iff)
+        hence "a * u x \<le> f i x" by auto
+        thus ?thesis using `x \<in> space M` by auto
+      qed
     qed
-  qed auto
-  note measure_conv = measure_up[OF Int[OF u B] this]
+    then show "?thesis i" using continuity_from_below[OF 1 2] by simp
+  qed
 
   have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
     unfolding simple_integral_indicator[OF B `simple_function M u`]
-  proof (subst SUPR_pextreal_setsum, safe)
+  proof (subst SUPR_extreal_setsum, safe)
     fix x n assume "x \<in> space M"
-    have "\<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f n x})
-      \<le> \<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f (Suc n) x})"
-      using B_mono Int[OF u[OF `x \<in> space M`] B] by (auto intro!: measure_mono)
-    thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n)
-            \<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
-      by (auto intro: mult_left_mono)
+    with u_range show "incseq (\<lambda>i. u x * \<mu> (?B' (u x) i))" "\<And>i. 0 \<le> u x * \<mu> (?B' (u x) i)"
+      using B_mono B_u by (auto intro!: measure_mono extreal_mult_left_mono incseq_SucI simp: extreal_zero_le_0_iff)
   next
-    show "integral\<^isup>S M u =
-      (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (u -` {i} \<inter> space M \<inter> ?B n))"
-      using measure_conv unfolding simple_integral_def isoton_def
-      by (auto intro!: setsum_cong simp: pextreal_SUP_cmult)
+    show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
+      using measure_conv u_range B_u unfolding simple_integral_def
+      by (auto intro!: setsum_cong SUPR_extreal_cmult[symmetric])
   qed
   moreover
   have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
-    unfolding pextreal_SUP_cmult[symmetric]
+    apply (subst SUPR_extreal_cmult[symmetric])
   proof (safe intro!: SUP_mono bexI)
     fix i
     have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
-      using B `simple_function M u`
-      by (subst simple_integral_mult[symmetric]) (auto simp: field_simps)
+      using B `simple_function M u` u_range
+      by (subst simple_integral_mult) (auto split: split_indicator)
     also have "\<dots> \<le> integral\<^isup>P M (f i)"
     proof -
-      have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto
-      hence *: "simple_function M (\<lambda>x. a * ?uB i x)" using B assms(3)
-        by (auto intro!: simple_integral_mono)
-      show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric]
-        by (auto intro!: positive_integral_mono simp: indicator_def)
+      have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B `0 < a` u(1) by auto
+      show ?thesis using f(3) * u_range `0 < a`
+        by (subst positive_integral_eq_simple_integral[symmetric])
+           (auto intro!: positive_integral_mono split: split_indicator)
     qed
     finally show "a * integral\<^isup>S M (?uB i) \<le> integral\<^isup>P M (f i)"
       by auto
-  qed simp
+  next
+    fix i show "0 \<le> \<integral>\<^isup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
+      by (intro simple_integral_positive) (auto split: split_indicator)
+  qed (insert `0 < a`, auto)
   ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
 qed
 
+lemma (in measure_space) incseq_positive_integral:
+  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^isup>P M (f i))"
+proof -
+  have "\<And>i x. f i x \<le> f (Suc i) x"
+    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
+  then show ?thesis
+    by (auto intro!: incseq_SucI positive_integral_mono)
+qed
+
 text {* Beppo-Levi monotone convergence theorem *}
-lemma (in measure_space) positive_integral_isoton:
-  assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>i. integral\<^isup>P M (f i)) \<up> integral\<^isup>P M u"
-  unfolding isoton_def
-proof safe
-  fix i show "integral\<^isup>P M (f i) \<le> integral\<^isup>P M (f (Suc i))"
-    apply (rule positive_integral_mono)
-    using `f \<up> u` unfolding isoton_def le_fun_def by auto
+lemma (in measure_space) positive_integral_monotone_convergence_SUP:
+  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+proof (rule antisym)
+  show "(SUP j. integral\<^isup>P M (f j)) \<le> (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
+    by (auto intro!: SUP_leI le_SUPI positive_integral_mono)
 next
-  have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
-  show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u"
-  proof (rule antisym)
-    from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def]
-    show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M u"
-      by (auto intro!: SUP_leI positive_integral_mono)
-  next
-    show "integral\<^isup>P M u \<le> (SUP i. integral\<^isup>P M (f i))"
-      unfolding positive_integral_alt[of u]
-      by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms])
+  show "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^isup>P M (f j))"
+    unfolding positive_integral_def_finite[of "\<lambda>x. SUP i. f i x"]
+  proof (safe intro!: SUP_leI)
+    fix g assume g: "simple_function M g"
+      and "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
+    moreover then have "\<And>x. 0 \<le> (SUP i. f i x)" and g': "g`space M \<subseteq> {0..<\<infinity>}"
+      using f by (auto intro!: le_SUPI2)
+    ultimately show "integral\<^isup>S M g \<le> (SUP j. integral\<^isup>P M (f j))"
+      by (intro  positive_integral_SUP_approx[OF f g _ g'])
+         (auto simp: le_fun_def max_def SUPR_apply)
   qed
 qed
 
-lemma (in measure_space) positive_integral_monotone_convergence_SUP:
-  assumes "\<And>i x. x \<in> space M \<Longrightarrow> f i x \<le> f (Suc i) x"
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(SUP i. integral\<^isup>P M (f i)) = (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
-    (is "_ = integral\<^isup>P M ?u")
+lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE:
+  assumes f: "\<And>i. AE x. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
 proof -
-  show ?thesis
-  proof (rule antisym)
-    show "(SUP j. integral\<^isup>P M (f j)) \<le> integral\<^isup>P M ?u"
-      by (auto intro!: SUP_leI positive_integral_mono le_SUPI)
-  next
-    def rf \<equiv> "\<lambda>i. \<lambda>x\<in>space M. f i x" and ru \<equiv> "\<lambda>x\<in>space M. ?u x"
-    have "\<And>i. rf i \<in> borel_measurable M" unfolding rf_def
-      using assms by (simp cong: measurable_cong)
-    moreover have iso: "rf \<up> ru" using assms unfolding rf_def ru_def
-      unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply
-      using SUP_const[OF UNIV_not_empty]
-      by (auto simp: restrict_def le_fun_def fun_eq_iff)
-    ultimately have "integral\<^isup>P M ru \<le> (SUP i. integral\<^isup>P M (rf i))"
-      unfolding positive_integral_alt[of ru]
-      by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx)
-    then show "integral\<^isup>P M ?u \<le> (SUP i. integral\<^isup>P M (f i))"
-      unfolding ru_def rf_def by (simp cong: positive_integral_cong)
+  from f have "AE x. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
+    by (simp add: AE_all_countable)
+  from this[THEN AE_E] guess N . note N = this
+  let "?f i x" = "if x \<in> space M - N then f i x else 0"
+  have f_eq: "AE x. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N])
+  then have "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP i. ?f i x) \<partial>M)"
+    by (auto intro!: positive_integral_cong_AE)
+  also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. ?f i x \<partial>M))"
+  proof (rule positive_integral_monotone_convergence_SUP)
+    show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
+    { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
+        using f N(3) by (intro measurable_If_set) auto
+      fix x show "0 \<le> ?f i x"
+        using N(1) by auto }
   qed
+  also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. f i x \<partial>M))"
+    using f_eq by (force intro!: arg_cong[where f="SUPR UNIV"] positive_integral_cong_AE ext)
+  finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq:
+  assumes f: "incseq f" "\<And>i. AE x. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+  using f[unfolded incseq_Suc_iff le_fun_def]
+  by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel])
+     auto
+
+lemma (in measure_space) positive_integral_monotone_convergence_simple:
+  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
+  shows "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+  using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1)
+    f(3)[THEN borel_measurable_simple_function] f(2)]
+  by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPR UNIV"] ext)
+
+lemma positive_integral_max_0:
+  "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M) = integral\<^isup>P M f"
+  by (simp add: le_fun_def positive_integral_def)
+
+lemma (in measure_space) positive_integral_cong_pos:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
+  shows "integral\<^isup>P M f = integral\<^isup>P M g"
+proof -
+  have "integral\<^isup>P M (\<lambda>x. max 0 (f x)) = integral\<^isup>P M (\<lambda>x. max 0 (g x))"
+  proof (intro positive_integral_cong)
+    fix x assume "x \<in> space M"
+    from assms[OF this] show "max 0 (f x) = max 0 (g x)"
+      by (auto split: split_max)
+  qed
+  then show ?thesis by (simp add: positive_integral_max_0)
 qed
 
 lemma (in measure_space) SUP_simple_integral_sequences:
-  assumes f: "f \<up> u" "\<And>i. simple_function M (f i)"
-  and g: "g \<up> u" "\<And>i. simple_function M (g i)"
+  assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
+  and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
+  and eq: "AE x. (SUP i. f i x) = (SUP i. g i x)"
   shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))"
     (is "SUPR _ ?F = SUPR _ ?G")
 proof -
-  have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))"
-    using assms by (simp add: positive_integral_eq_simple_integral)
-  also have "\<dots> = integral\<^isup>P M u"
-    using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]]
-    unfolding isoton_def by simp
-  also have "\<dots> = (SUP i. integral\<^isup>P M (g i))"
-    using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]]
-    unfolding isoton_def by simp
+  have "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+    using f by (rule positive_integral_monotone_convergence_simple)
+  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. g i x) \<partial>M)"
+    unfolding eq[THEN positive_integral_cong_AE] ..
   also have "\<dots> = (SUP i. ?G i)"
-    using assms by (simp add: positive_integral_eq_simple_integral)
-  finally show ?thesis .
+    using g by (rule positive_integral_monotone_convergence_simple[symmetric])
+  finally show ?thesis by simp
 qed
 
 lemma (in measure_space) positive_integral_const[simp]:
-  "(\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
+  "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
   by (subst positive_integral_eq_simple_integral) auto
 
-lemma (in measure_space) positive_integral_isoton_simple:
-  assumes "f \<up> u" and e: "\<And>i. simple_function M (f i)"
-  shows "(\<lambda>i. integral\<^isup>S M (f i)) \<up> integral\<^isup>P M u"
-  using positive_integral_isoton[OF `f \<up> u` e(1)[THEN borel_measurable_simple_function]]
-  unfolding positive_integral_eq_simple_integral[OF e] .
-
-lemma measure_preservingD2:
-  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
-  unfolding measure_preserving_def by auto
-
 lemma (in measure_space) positive_integral_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" and f: "f \<in> borel_measurable M'"
+  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
+  and f: "f \<in> borel_measurable M'"
   shows "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
 proof -
   interpret T: measure_space M' by (rule measure_space_vimage[OF T])
-  obtain f' where f': "f' \<up> f" "\<And>i. simple_function M' (f' i)"
-    using T.borel_measurable_implies_simple_function_sequence[OF f] by blast
-  then have f: "(\<lambda>i x. f' i (T x)) \<up> (\<lambda>x. f (T x))" "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
-    using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)]] unfolding isoton_fun_expand by auto
+  from T.borel_measurable_implies_simple_function_sequence'[OF f]
+  guess f' . note f' = this
+  let "?f i x" = "f' i (T x)"
+  have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
+  have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
+    using f'(4) .
+  have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
+    using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(1)] .
   show "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
-    using positive_integral_isoton_simple[OF f]
-    using T.positive_integral_isoton_simple[OF f']
-    by (simp add: simple_integral_vimage[OF T f'(2)] isoton_def)
+    using
+      T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)]
+      positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
+    by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f')
 qed
 
 lemma (in measure_space) positive_integral_linear:
-  assumes f: "f \<in> borel_measurable M"
-  and g: "g \<in> borel_measurable M"
+  assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
+  and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
   shows "(\<integral>\<^isup>+ x. a * f x + g x \<partial>M) = a * integral\<^isup>P M f + integral\<^isup>P M g"
     (is "integral\<^isup>P M ?L = _")
 proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f] guess u .
-  note u = this positive_integral_isoton_simple[OF this(1-2)]
-  from borel_measurable_implies_simple_function_sequence'[OF g] guess v .
-  note v = this positive_integral_isoton_simple[OF this(1-2)]
+  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
+  note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
+  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
+  note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
   let "?L' i x" = "a * u i x + v i x"
 
-  have "?L \<in> borel_measurable M"
-    using assms by simp
+  have "?L \<in> borel_measurable M" using assms by auto
   from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
-  note positive_integral_isoton_simple[OF this(1-2)] and l = this
-  moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
-  proof (rule SUP_simple_integral_sequences[OF l(1-2)])
-    show "?L' \<up> ?L" "\<And>i. simple_function M (?L' i)"
-      using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right)
+  note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
+
+  have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
+    using u v `0 \<le> a`
+    by (auto simp: incseq_Suc_iff le_fun_def
+             intro!: add_mono extreal_mult_left_mono simple_integral_mono)
+  have pos: "\<And>i. 0 \<le> integral\<^isup>S M (u i)" "\<And>i. 0 \<le> integral\<^isup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^isup>S M (u i)"
+    using u v `0 \<le> a` by (auto simp: simple_integral_positive)
+  { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
+      by (auto split: split_if_asm) }
+  note not_MInf = this
+
+  have l': "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
+  proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
+    show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
+      using u v  `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
+      by (auto intro!: add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg)
+    { fix x
+      { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
+          by auto }
+      then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
+        using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
+        by (subst SUPR_extreal_cmult[symmetric, OF u(6) `0 \<le> a`])
+           (auto intro!: SUPR_extreal_add
+                 simp: incseq_Suc_iff le_fun_def add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) }
+    then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
+      unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
+      by (intro AE_I2) (auto split: split_max simp add: extreal_add_nonneg_nonneg)
   qed
-  moreover from u v have L'_isoton:
-      "(\<lambda>i. integral\<^isup>S M (?L' i)) \<up> a * integral\<^isup>P M f + integral\<^isup>P M g"
-    by (simp add: isoton_add isoton_cmult_right)
-  ultimately show ?thesis by (simp add: isoton_def)
+  also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
+    using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
+  finally have "(\<integral>\<^isup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^isup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+x. max 0 (g x) \<partial>M)"
+    unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
+    unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
+    apply (subst SUPR_extreal_cmult[symmetric, OF pos(1) `0 \<le> a`])
+    apply (subst SUPR_extreal_add[symmetric, OF inc not_MInf]) .
+  then show ?thesis by (simp add: positive_integral_max_0)
 qed
 
 lemma (in measure_space) positive_integral_cmult:
-  assumes "f \<in> borel_measurable M"
+  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
-  using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto
+proof -
+  have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
+    by (auto split: split_max simp: extreal_zero_le_0_iff)
+  have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
+    by (simp add: positive_integral_max_0)
+  then show ?thesis
+    using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" "\<lambda>x. 0"] f
+    by (auto simp: positive_integral_max_0)
+qed
 
 lemma (in measure_space) positive_integral_multc:
-  assumes "f \<in> borel_measurable M"
+  assumes "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
   unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
 
@@ -1260,143 +1289,172 @@
      (auto simp: simple_function_indicator simple_integral_indicator)
 
 lemma (in measure_space) positive_integral_cmult_indicator:
-  "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
+  "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
   by (subst positive_integral_eq_simple_integral)
      (auto simp: simple_function_indicator simple_integral_indicator)
 
 lemma (in measure_space) positive_integral_add:
-  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
+  and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
   shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
-  using positive_integral_linear[OF assms, of 1] by simp
+proof -
+  have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
+    using assms by (auto split: split_max simp: extreal_add_nonneg_nonneg)
+  have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
+    by (simp add: positive_integral_max_0)
+  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
+    unfolding ae[THEN positive_integral_cong_AE] ..
+  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+ x. max 0 (g x) \<partial>M)"
+    using positive_integral_linear[of "\<lambda>x. max 0 (f x)" 1 "\<lambda>x. max 0 (g x)"] f g
+    by auto
+  finally show ?thesis
+    by (simp add: positive_integral_max_0)
+qed
 
 lemma (in measure_space) positive_integral_setsum:
-  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
+  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x. 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>P M (f i))"
 proof cases
-  assume "finite P"
-  from this assms show ?thesis
+  assume f: "finite P"
+  from assms have "AE x. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
+  from f this assms(1) show ?thesis
   proof induct
     case (insert i P)
-    have "f i \<in> borel_measurable M"
-      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M"
-      using insert by (auto intro!: borel_measurable_pextreal_setsum)
+    then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
+      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
+      by (auto intro!: borel_measurable_extreal_setsum setsum_nonneg)
     from positive_integral_add[OF this]
     show ?case using insert by auto
   qed simp
 qed simp
 
+lemma (in measure_space) positive_integral_Markov_inequality:
+  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
+  shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+    (is "\<mu> ?A \<le> _ * ?PI")
+proof -
+  have "?A \<in> sets M"
+    using `A \<in> sets M` u by auto
+  hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
+    using positive_integral_indicator by simp
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
+    by (auto intro!: positive_integral_mono_AE
+      simp: indicator_def extreal_zero_le_0_iff)
+  also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+    using assms
+    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: extreal_zero_le_0_iff)
+  finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_noteq_infinite:
+  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+  and "integral\<^isup>P M g \<noteq> \<infinity>"
+  shows "AE x. g x \<noteq> \<infinity>"
+proof (rule ccontr)
+  assume c: "\<not> (AE x. g x \<noteq> \<infinity>)"
+  have "\<mu> {x\<in>space M. g x = \<infinity>} \<noteq> 0"
+    using c g by (simp add: AE_iff_null_set)
+  moreover have "0 \<le> \<mu> {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
+  ultimately have "0 < \<mu> {x\<in>space M. g x = \<infinity>}" by auto
+  then have "\<infinity> = \<infinity> * \<mu> {x\<in>space M. g x = \<infinity>}" by auto
+  also have "\<dots> \<le> (\<integral>\<^isup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
+    using g by (subst positive_integral_cmult_indicator) auto
+  also have "\<dots> \<le> integral\<^isup>P M g"
+    using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def)
+  finally show False using `integral\<^isup>P M g \<noteq> \<infinity>` by auto
+qed
+
 lemma (in measure_space) positive_integral_diff:
-  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
-  and fin: "integral\<^isup>P M g \<noteq> \<omega>"
-  and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
+  assumes f: "f \<in> borel_measurable M"
+  and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+  and fin: "integral\<^isup>P M g \<noteq> \<infinity>"
+  and mono: "AE x. g x \<le> f x"
   shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
 proof -
-  have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-    using f g by (rule borel_measurable_pextreal_diff)
-  have "(\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g = integral\<^isup>P M f"
-    unfolding positive_integral_add[OF borel g, symmetric]
-  proof (rule positive_integral_cong)
-    fix x assume "x \<in> space M"
-    from mono[OF this] show "f x - g x + g x = f x"
-      by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
-  qed
-  with mono show ?thesis
-    by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
+  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
+    using assms by (auto intro: extreal_diff_positive)
+  have pos_f: "AE x. 0 \<le> f x" using mono g by auto
+  { fix a b :: extreal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
+      by (cases rule: extreal2_cases[of a b]) auto }
+  note * = this
+  then have "AE x. f x = f x - g x + g x"
+    using mono positive_integral_noteq_infinite[OF g fin] assms by auto
+  then have **: "integral\<^isup>P M f = (\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g"
+    unfolding positive_integral_add[OF diff g, symmetric]
+    by (rule positive_integral_cong_AE)
+  show ?thesis unfolding **
+    using fin positive_integral_positive[of g]
+    by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
 qed
 
-lemma (in measure_space) positive_integral_psuminf:
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<integral>\<^isup>+ x. (\<Sum>\<^isub>\<infinity> i. f i x) \<partial>M) = (\<Sum>\<^isub>\<infinity> i. integral\<^isup>P M (f i))"
+lemma (in measure_space) positive_integral_suminf:
+  assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> f i x"
+  shows "(\<integral>\<^isup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^isup>P M (f i))"
 proof -
-  have "(\<lambda>i. (\<integral>\<^isup>+x. (\<Sum>i<i. f i x) \<partial>M)) \<up> (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>i. f i x) \<partial>M)"
-    by (rule positive_integral_isoton)
-       (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono
-                     arg_cong[where f=Sup]
-             simp: isoton_def le_fun_def psuminf_def fun_eq_iff SUPR_def Sup_fun_def)
-  thus ?thesis
-    by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms])
+  have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
+    using assms by (auto simp: AE_all_countable)
+  have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
+    using positive_integral_positive by (rule suminf_extreal_eq_SUPR)
+  also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
+    unfolding positive_integral_setsum[OF f] ..
+  also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
+    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
+       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
+  also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
+    by (intro positive_integral_cong_AE) (auto simp: suminf_extreal_eq_SUPR)
+  finally show ?thesis by simp
 qed
 
 text {* Fatou's lemma: convergence theorem on limes inferior *}
 lemma (in measure_space) positive_integral_lim_INF:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pextreal"
-  assumes "\<And>i. u i \<in> borel_measurable M"
-  shows "(\<integral>\<^isup>+ x. (SUP n. INF m. u (m + n) x) \<partial>M) \<le>
-    (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
+  shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
 proof -
-  have "(\<integral>\<^isup>+x. (SUP n. INF m. u (m + n) x) \<partial>M)
-      = (SUP n. (\<integral>\<^isup>+x. (INF m. u (m + n) x) \<partial>M))"
-    using assms
-    by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono)
-       (auto simp del: add_Suc simp add: add_Suc[symmetric])
-  also have "\<dots> \<le> (SUP n. INF m. integral\<^isup>P M (u (m + n)))"
-    by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI)
+  have pos: "AE x. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
+  have "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
+    (SUP n. \<integral>\<^isup>+ x. (INF i:{n..}. u i x) \<partial>M)"
+    unfolding liminf_SUPR_INFI using pos u
+    by (intro positive_integral_monotone_convergence_SUP_AE)
+       (elim AE_mp, auto intro!: AE_I2 intro: le_INFI INF_subset)
+  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
+    unfolding liminf_SUPR_INFI
+    by (auto intro!: SUP_mono exI le_INFI positive_integral_mono INF_leI)
   finally show ?thesis .
 qed
 
 lemma (in measure_space) measure_space_density:
-  assumes borel: "u \<in> borel_measurable M"
+  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x"
     and M'[simp]: "M' = (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)\<rparr>)"
   shows "measure_space M'"
 proof -
   interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto
   show ?thesis
   proof
-    show "measure M' {} = 0" unfolding M' by simp
+    have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
+      using u by (auto simp: extreal_zero_le_0_iff)
+    then show "positive M' (measure M')" unfolding M'
+      using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
     show "countably_additive M' (measure M')"
     proof (intro countably_additiveI)
       fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M'"
-      then have "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
-        using borel by (auto intro: borel_measurable_indicator)
-      moreover assume "disjoint_family A"
-      note psuminf_indicator[OF this]
-      ultimately show "(\<Sum>\<^isub>\<infinity>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
-        by (simp add: positive_integral_psuminf[symmetric])
+      then have *: "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
+        using u by (auto intro: borel_measurable_indicator)
+      assume disj: "disjoint_family A"
+      have "(\<Sum>n. measure M' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. u x * indicator (A n) x) \<partial>M)"
+        unfolding M' using u(1) *
+        by (simp add: positive_integral_suminf[OF _ pos, symmetric])
+      also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
+        by (intro positive_integral_cong_AE)
+           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_extreal)
+      also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
+        unfolding suminf_indicator[OF disj] ..
+      finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
+        unfolding M' by simp
     qed
   qed
 qed
 
-lemma (in measure_space) positive_integral_translated_density:
-  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-    and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
-  shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
-proof -
-  from measure_space_density[OF assms(1) M']
-  interpret T: measure_space M' .
-  have borel[simp]:
-    "borel_measurable M' = borel_measurable M"
-    "simple_function M' = simple_function M"
-    unfolding measurable_def simple_function_def_raw by (auto simp: M')
-  from borel_measurable_implies_simple_function_sequence[OF assms(2)]
-  obtain G where G: "\<And>i. simple_function M (G i)" "G \<up> g" by blast
-  note G_borel = borel_measurable_simple_function[OF this(1)]
-  from T.positive_integral_isoton[unfolded borel, OF `G \<up> g` G_borel]
-  have *: "(\<lambda>i. integral\<^isup>P M' (G i)) \<up> integral\<^isup>P M' g" .
-  { fix i
-    have [simp]: "finite (G i ` space M)"
-      using G(1) unfolding simple_function_def by auto
-    have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
-      using G T.positive_integral_eq_simple_integral by simp
-    also have "\<dots> = (\<integral>\<^isup>+x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x) \<partial>M)"
-      apply (simp add: simple_integral_def M')
-      apply (subst positive_integral_cmult[symmetric])
-      using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
-      apply (subst positive_integral_setsum[symmetric])
-      using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage)
-      by (simp add: setsum_right_distrib field_simps)
-    also have "\<dots> = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
-      by (auto intro!: positive_integral_cong
-               simp: indicator_def if_distrib setsum_cases)
-    finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)" . }
-  with * have eq_Tg: "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> integral\<^isup>P M' g" by simp
-  from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
-    unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
-  then have "(\<lambda>i. (\<integral>\<^isup>+x. f x * G i x \<partial>M)) \<up> (\<integral>\<^isup>+x. f x * g x \<partial>M)"
-    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times)
-  with eq_Tg show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)"
-    unfolding isoton_def by simp
-qed
-
 lemma (in measure_space) positive_integral_null_set:
   assumes "N \<in> null_sets" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
 proof -
@@ -1410,144 +1468,199 @@
   then show ?thesis by simp
 qed
 
-lemma (in measure_space) positive_integral_Markov_inequality:
-  assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>"
-  shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
-    (is "\<mu> ?A \<le> _ * ?PI")
+lemma (in measure_space) positive_integral_translated_density:
+  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
+  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+    and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
+  shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
 proof -
-  have "?A \<in> sets M"
-    using `A \<in> sets M` borel by auto
-  hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
-    using positive_integral_indicator by simp
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)"
-  proof (rule positive_integral_mono)
-    fix x assume "x \<in> space M"
-    show "indicator ?A x \<le> c * (u x * indicator A x)"
-      by (cases "x \<in> ?A") auto
-  qed
-  also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
-    using assms
-    by (auto intro!: positive_integral_cmult borel_measurable_indicator)
-  finally show ?thesis .
+  from measure_space_density[OF f M']
+  interpret T: measure_space M' .
+  have borel[simp]:
+    "borel_measurable M' = borel_measurable M"
+    "simple_function M' = simple_function M"
+    unfolding measurable_def simple_function_def_raw by (auto simp: M')
+  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
+  note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
+  note G'(2)[simp]
+  { fix P have "AE x. P x \<Longrightarrow> AE x in M'. P x"
+      using positive_integral_null_set[of _ f]
+      unfolding T.almost_everywhere_def almost_everywhere_def
+      by (auto simp: M') }
+  note ac = this
+  from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x"
+    by (auto intro!: ac split: split_max)
+  { fix i
+    let "?I y x" = "indicator (G i -` {y} \<inter> space M) x"
+    { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
+      then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
+      from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
+        by (subst setsum_extreal_right_distrib) (auto simp: ac_simps)
+      also have "\<dots> = f x * G i x"
+        by (simp add: indicator_def if_distrib setsum_cases)
+      finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
+    note to_singleton = this
+    have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
+      using G T.positive_integral_eq_simple_integral by simp
+    also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
+      unfolding simple_integral_def M' by simp
+    also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
+      using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
+    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
+      using f G' G by (auto intro!: positive_integral_setsum[symmetric])
+    finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
+      using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
+  note [simp] = this
+  have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G
+    using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`]
+    by (simp cong: T.positive_integral_cong_AE)
+  also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
+  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
+    using f G' G(2)[THEN incseq_SucD] G
+    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
+       (auto simp: extreal_mult_left_mono le_fun_def extreal_zero_le_0_iff)
+  also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
+    by (intro positive_integral_cong_AE)
+       (auto simp add: SUPR_extreal_cmult split: split_max)
+  finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
 qed
 
 lemma (in measure_space) positive_integral_0_iff:
-  assumes borel: "u \<in> borel_measurable M"
+  assumes u: "u \<in> borel_measurable M" and pos: "AE x. 0 \<le> u x"
   shows "integral\<^isup>P M u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
     (is "_ \<longleftrightarrow> \<mu> ?A = 0")
 proof -
-  have A: "?A \<in> sets M" using borel by auto
-  have u: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
+  have u_eq: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
     by (auto intro!: positive_integral_cong simp: indicator_def)
-
   show ?thesis
   proof
     assume "\<mu> ?A = 0"
-    hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
-    from positive_integral_null_set[OF this]
-    have "0 = (\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M)" by simp
-    thus "integral\<^isup>P M u = 0" unfolding u by simp
+    with positive_integral_null_set[of ?A u] u
+    show "integral\<^isup>P M u = 0" by (simp add: u_eq)
   next
+    { fix r :: extreal and n :: nat assume gt_1: "1 \<le> real n * r"
+      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_extreal_def)
+      then have "0 \<le> r" by (auto simp add: extreal_zero_less_0_iff) }
+    note gt_1 = this
     assume *: "integral\<^isup>P M u = 0"
-    let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
+    let "?M n" = "{x \<in> space M. 1 \<le> real (n::nat) * u x}"
     have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
     proof -
-      { fix n
-        from positive_integral_Markov_inequality[OF borel `?A \<in> sets M`, of "of_nat n"]
-        have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp }
+      { fix n :: nat
+        from positive_integral_Markov_inequality[OF u pos, of ?A "extreal (real n)"]
+        have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
+        moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
+        ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
       thus ?thesis by simp
     qed
     also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
     proof (safe intro!: continuity_from_below)
       fix n show "?M n \<inter> ?A \<in> sets M"
-        using borel by (auto intro!: Int)
+        using u by (auto intro!: Int)
     next
-      fix n x assume "1 \<le> of_nat n * u x"
-      also have "\<dots> \<le> of_nat (Suc n) * u x"
-        by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel)
-      finally show "1 \<le> of_nat (Suc n) * u x" .
+      show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
+      proof (safe intro!: incseq_SucI)
+        fix n :: nat and x
+        assume *: "1 \<le> real n * u x"
+        also from gt_1[OF this] have "real n * u x \<le> real (Suc n) * u x"
+          using `0 \<le> u x` by (auto intro!: extreal_mult_right_mono)
+        finally show "1 \<le> real (Suc n) * u x" by auto
+      qed
     qed
-    also have "\<dots> = \<mu> ?A"
-    proof (safe intro!: arg_cong[where f="\<mu>"])
-      fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M"
+    also have "\<dots> = \<mu> {x\<in>space M. 0 < u x}"
+    proof (safe intro!: arg_cong[where f="\<mu>"] dest!: gt_1)
+      fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
       show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
       proof (cases "u x")
-        case (preal r)
-        obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat ..
-        hence "1 / r * r \<le> of_nat j * r" using preal unfolding mult_le_cancel_right by auto
-        hence "1 \<le> of_nat j * r" using preal `u x \<noteq> 0` by auto
-        thus ?thesis using `u x \<noteq> 0` preal by (auto simp: real_of_nat_def[symmetric])
-      qed auto
-    qed
-    finally show "\<mu> ?A = 0" by simp
+        case (real r) with `0 < u x` have "0 < r" by auto
+        obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
+        hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
+        hence "1 \<le> real j * r" using real `0 < r` by auto
+        thus ?thesis using `0 < r` real by (auto simp: one_extreal_def)
+      qed (insert `0 < u x`, auto)
+    qed auto
+    finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
+    moreover
+    from pos have "AE x. \<not> (u x < 0)" by auto
+    then have "\<mu> {x\<in>space M. u x < 0} = 0"
+      using AE_iff_null_set u by auto
+    moreover have "\<mu> {x\<in>space M. u x \<noteq> 0} = \<mu> {x\<in>space M. u x < 0} + \<mu> {x\<in>space M. 0 < u x}"
+      using u by (subst measure_additive) (auto intro!: arg_cong[where f=\<mu>])
+    ultimately show "\<mu> ?A = 0" by simp
   qed
 qed
 
 lemma (in measure_space) positive_integral_0_iff_AE:
   assumes u: "u \<in> borel_measurable M"
-  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x = 0)"
+  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x \<le> 0)"
 proof -
-  have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
+  have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
     using u by auto
-  then show ?thesis
-    using positive_integral_0_iff[OF u] AE_iff_null_set[OF sets] by auto
+  from positive_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
+  have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. max 0 (u x) = 0)"
+    unfolding positive_integral_max_0
+    using AE_iff_null_set[OF sets] u by auto
+  also have "\<dots> \<longleftrightarrow> (AE x. u x \<le> 0)" by (auto split: split_max)
+  finally show ?thesis .
 qed
 
 lemma (in measure_space) positive_integral_restricted:
-  assumes "A \<in> sets M"
+  assumes A: "A \<in> sets M"
   shows "integral\<^isup>P (restricted_space A) f = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
     (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f")
 proof -
-  have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \<in> sets M`])
-  then interpret R: measure_space ?R .
-  have saR: "sigma_algebra ?R" by fact
-  have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f"
-    by (intro R.positive_integral_cong) auto
+  interpret R: measure_space ?R
+    by (rule restricted_measure_space) fact
+  let "?I g x" = "g x * indicator A x :: extreal"
   show ?thesis
-    unfolding * positive_integral_def
-    unfolding simple_function_restricted[OF `A \<in> sets M`]
-    apply (simp add: SUPR_def)
-    apply (rule arg_cong[where f=Sup])
-  proof (auto simp add: image_iff simple_integral_restricted[OF `A \<in> sets M`])
-    fix g assume "simple_function M (\<lambda>x. g x * indicator A x)"
-      "g \<le> f"
-    then show "\<exists>x. simple_function M x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and>
-      (\<integral>\<^isup>Sx. g x * indicator A x \<partial>M) = integral\<^isup>S M x"
-      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
-      by (auto simp: indicator_def le_fun_def)
+    unfolding positive_integral_def
+    unfolding simple_function_restricted[OF A]
+    unfolding AE_restricted[OF A]
+  proof (safe intro!: SUPR_eq)
+    fix g assume g: "simple_function M (?I g)" and le: "g \<le> max 0 \<circ> f"
+    show "\<exists>j\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> ?I f}.
+      integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M j"
+    proof (safe intro!: bexI[of _ "?I g"])
+      show "integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M (?I g)"
+        using g A by (simp add: simple_integral_restricted)
+      show "?I g \<le> max 0 \<circ> ?I f"
+        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+    qed fact
   next
-    fix g assume g: "simple_function M g" "g \<le> (\<lambda>x. f x * indicator A x)"
-    then have *: "(\<lambda>x. g x * indicator A x) = g"
-      "\<And>x. g x * indicator A x = g x"
-      "\<And>x. g x \<le> f x"
-      by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm)
-    from g show "\<exists>x. simple_function M (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and>
-      integral\<^isup>S M g = integral\<^isup>S M (\<lambda>xa. x xa * indicator A xa)"
-      using `A \<in> sets M`[THEN sets_into_space]
-      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
-      by (fastsimp simp: le_fun_def *)
+    fix g assume g: "simple_function M g" and le: "g \<le> max 0 \<circ> ?I f"
+    show "\<exists>i\<in>{g. simple_function M (?I g) \<and> g \<le> max 0 \<circ> f}.
+      integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) i"
+    proof (safe intro!: bexI[of _ "?I g"])
+      show "?I g \<le> max 0 \<circ> f"
+        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+      from le have "\<And>x. g x \<le> ?I (?I g) x"
+        by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
+      then show "integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) (?I g)"
+        using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted)
+      show "simple_function M (?I (?I g))" using g A by auto
+    qed
   qed
 qed
 
 lemma (in measure_space) positive_integral_subalgebra:
-  assumes borel: "f \<in> borel_measurable N"
+  assumes f: "f \<in> borel_measurable N" "AE x in N. 0 \<le> f x"
   and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
   and sa: "sigma_algebra N"
   shows "integral\<^isup>P N f = integral\<^isup>P M f"
 proof -
   interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
-  from N.borel_measurable_implies_simple_function_sequence[OF borel]
-  obtain fs where Nsf: "\<And>i. simple_function N (fs i)" and "fs \<up> f" by blast
-  then have sf: "\<And>i. simple_function M (fs i)"
-    using simple_function_subalgebra[OF _ N(1,2)] by blast
-  from N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf]
+  from N.borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
+  note sf = simple_function_subalgebra[OF fs(1) N(1,2)]
+  from N.positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric]
   have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * N.\<mu> (fs i -` {x} \<inter> space M))"
-    unfolding isoton_def simple_integral_def `space N = space M` by simp
+    unfolding fs(4) positive_integral_max_0
+    unfolding simple_integral_def `space N = space M` by simp
   also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * \<mu> (fs i -` {x} \<inter> space M))"
-    using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto
+    using N N.simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
   also have "\<dots> = integral\<^isup>P M f"
-    using positive_integral_isoton_simple[OF `fs \<up> f` sf]
-    unfolding isoton_def simple_integral_def `space N = space M` by simp
+    using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric]
+    unfolding fs(4) positive_integral_max_0
+    unfolding simple_integral_def `space N = space M` by simp
   finally show ?thesis .
 qed
 
@@ -1555,16 +1668,15 @@
 
 definition integrable where
   "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
-    (\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega> \<and>
-    (\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
+    (\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
 
 lemma integrableD[dest]:
   assumes "integrable M f"
-  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) \<noteq> \<omega>" "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) \<noteq> \<omega>"
+  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
   using assms unfolding integrable_def by auto
 
 definition lebesgue_integral_def:
-  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. Real (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. Real (- f x) \<partial>M))"
+  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. extreal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. extreal (- f x) \<partial>M))"
 
 syntax
   "_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
@@ -1572,6 +1684,17 @@
 translations
   "\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
 
+lemma (in measure_space) integrableE:
+  assumes "integrable M f"
+  obtains r q where
+    "(\<integral>\<^isup>+x. extreal (f x)\<partial>M) = extreal r"
+    "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M) = extreal q"
+    "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
+  using assms unfolding integrable_def lebesgue_integral_def
+  using positive_integral_positive[of "\<lambda>x. extreal (f x)"]
+  using positive_integral_positive[of "\<lambda>x. extreal (-f x)"]
+  by (cases rule: extreal2_cases[of "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. extreal (f x)\<partial>M)"]) auto
+
 lemma (in measure_space) integral_cong:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
   shows "integral\<^isup>L M f = integral\<^isup>L M g"
@@ -1580,21 +1703,16 @@
 lemma (in measure_space) integral_cong_measure:
   assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
   shows "integral\<^isup>L N f = integral\<^isup>L M f"
-proof -
-  interpret v: measure_space N
-    by (rule measure_space_cong) fact+
-  show ?thesis
-    by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
-qed
+  by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
 
 lemma (in measure_space) integral_cong_AE:
   assumes cong: "AE x. f x = g x"
   shows "integral\<^isup>L M f = integral\<^isup>L M g"
 proof -
-  have "AE x. Real (f x) = Real (g x)" using cong by auto
-  moreover have "AE x. Real (- f x) = Real (- g x)" using cong by auto
-  ultimately show ?thesis
-    by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def)
+  have *: "AE x. extreal (f x) = extreal (g x)"
+    "AE x. extreal (- f x) = extreal (- g x)" using cong by auto
+  show ?thesis
+    unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
 qed
 
 lemma (in measure_space) integrable_cong:
@@ -1602,11 +1720,14 @@
   by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
 
 lemma (in measure_space) integral_eq_positive_integral:
-  assumes "\<And>x. 0 \<le> f x"
-  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+  assumes f: "\<And>x. 0 \<le> f x"
+  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
 proof -
-  have "\<And>x. Real (- f x) = 0" using assms by simp
-  thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def)
+  { fix x have "max 0 (extreal (- f x)) = 0" using f[of x] by (simp split: split_max) }
+  then have "0 = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)" by simp
+  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
+  finally show ?thesis
+    unfolding lebesgue_integral_def by simp
 qed
 
 lemma (in measure_space) integral_vimage:
@@ -1616,7 +1737,7 @@
 proof -
   interpret T: measure_space M' by (rule measure_space_vimage[OF T])
   from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable M'" "(\<lambda>x. Real (- f x)) \<in> borel_measurable M'"
+  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
     and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
     using f by (auto simp: comp_def)
   then show ?thesis
@@ -1631,7 +1752,7 @@
 proof -
   interpret T: measure_space M' by (rule measure_space_vimage[OF T])
   from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable M'" "(\<lambda>x. Real (- f x)) \<in> borel_measurable M'"
+  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
     and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
     using f by (auto simp: comp_def)
   then show ?thesis
@@ -1649,10 +1770,10 @@
   and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
   shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
 proof -
-  let "?f x" = "Real (f x)"
-  let "?mf x" = "Real (- f x)"
-  let "?u x" = "Real (u x)"
-  let "?v x" = "Real (v x)"
+  let "?f x" = "max 0 (extreal (f x))"
+  let "?mf x" = "max 0 (extreal (- f x))"
+  let "?u x" = "max 0 (extreal (u x))"
+  let "?v x" = "max 0 (extreal (v x))"
 
   from borel_measurable_diff[of u v] integrable
   have f_borel: "?f \<in> borel_measurable M" and
@@ -1662,73 +1783,62 @@
     "f \<in> borel_measurable M"
     by (auto simp: f_def[symmetric] integrable_def)
 
-  have "(\<integral>\<^isup>+ x. Real (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
-    using pos by (auto intro!: positive_integral_mono)
-  moreover have "(\<integral>\<^isup>+ x. Real (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
-    using pos by (auto intro!: positive_integral_mono)
+  have "(\<integral>\<^isup>+ x. extreal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
+    using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
+  moreover have "(\<integral>\<^isup>+ x. extreal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
+    using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
   ultimately show f: "integrable M f"
     using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
-    by (auto simp: integrable_def f_def)
-  hence mf: "integrable M (\<lambda>x. - f x)" ..
-
-  have *: "\<And>x. Real (- v x) = 0" "\<And>x. Real (- u x) = 0"
-    using pos by auto
+    by (auto simp: integrable_def f_def positive_integral_max_0)
 
   have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
-    unfolding f_def using pos by simp
-  hence "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
-  hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
+    unfolding f_def using pos by (simp split: split_max)
+  then have "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
+  then have "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
       real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)"
-    using positive_integral_add[OF u_borel mf_borel]
-    using positive_integral_add[OF v_borel f_borel]
+    using positive_integral_add[OF u_borel _ mf_borel]
+    using positive_integral_add[OF v_borel _ f_borel]
     by auto
   then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
-    using f mf `integrable M u` `integrable M v`
-    unfolding lebesgue_integral_def integrable_def *
-    by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u")
-       (auto simp add: field_simps)
+    unfolding positive_integral_max_0
+    unfolding pos[THEN integral_eq_positive_integral]
+    using integrable f by (auto elim!: integrableE)
 qed
 
 lemma (in measure_space) integral_linear:
   assumes "integrable M f" "integrable M g" and "0 \<le> a"
   shows "integrable M (\<lambda>t. a * f t + g t)"
-  and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
+  and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
 proof -
-  let ?PI = "integral\<^isup>P M"
-  let "?f x" = "Real (f x)"
-  let "?g x" = "Real (g x)"
-  let "?mf x" = "Real (- f x)"
-  let "?mg x" = "Real (- g x)"
+  let "?f x" = "max 0 (extreal (f x))"
+  let "?g x" = "max 0 (extreal (g x))"
+  let "?mf x" = "max 0 (extreal (- f x))"
+  let "?mg x" = "max 0 (extreal (- g x))"
   let "?p t" = "max 0 (a * f t) + max 0 (g t)"
   let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
 
-  have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M"
-    and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M"
-    and p: "?p \<in> borel_measurable M"
-    and n: "?n \<in> borel_measurable M"
-    using assms by (simp_all add: integrable_def)
+  from assms have linear:
+    "(\<integral>\<^isup>+ x. extreal a * ?f x + ?g x \<partial>M) = extreal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
+    "(\<integral>\<^isup>+ x. extreal a * ?mf x + ?mg x \<partial>M) = extreal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
+    by (auto intro!: positive_integral_linear simp: integrable_def)
 
-  have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x"
-          "\<And>x. Real (?n x) = Real a * ?mf x + ?mg x"
-          "\<And>x. Real (- ?p x) = 0"
-          "\<And>x. Real (- ?n x) = 0"
-    using `0 \<le> a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos)
-
-  note linear =
-    positive_integral_linear[OF pos]
-    positive_integral_linear[OF neg]
+  have *: "(\<integral>\<^isup>+x. extreal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- ?n x) \<partial>M) = 0"
+    using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
+  have **: "\<And>x. extreal a * ?f x + ?g x = max 0 (extreal (?p x))"
+           "\<And>x. extreal a * ?mf x + ?mg x = max 0 (extreal (?n x))"
+    using `0 \<le> a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff)
 
   have "integrable M ?p" "integrable M ?n"
       "\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t"
-    using assms p n unfolding integrable_def * linear by auto
+    using linear assms unfolding integrable_def ** *
+    by (auto simp: positive_integral_max_0)
   note diff = integral_of_positive_diff[OF this]
 
   show "integrable M (\<lambda>t. a * f t + g t)" by (rule diff)
-
-  from assms show "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g"
-    unfolding diff(2) unfolding lebesgue_integral_def * linear integrable_def
-    by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg")
-       (auto simp add: field_simps zero_le_mult_iff)
+  from assms linear show ?EQ
+    unfolding diff(2) ** positive_integral_max_0
+    unfolding lebesgue_integral_def *
+    by (auto elim!: integrableE simp: field_simps)
 qed
 
 lemma (in measure_space) integral_add[simp, intro]:
@@ -1772,13 +1882,13 @@
   and mono: "AE t. f t \<le> g t"
   shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
 proof -
-  have "AE x. Real (f x) \<le> Real (g x)"
+  have "AE x. extreal (f x) \<le> extreal (g x)"
     using mono by auto
-  moreover have "AE x. Real (- g x) \<le> Real (- f x)"
+  moreover have "AE x. extreal (- g x) \<le> extreal (- f x)"
     using mono by auto
   ultimately show ?thesis using fg
-    by (auto simp: lebesgue_integral_def integrable_def diff_minus
-             intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE)
+    by (auto intro!: add_mono positive_integral_mono_AE real_of_extreal_positive_mono
+             simp: positive_integral_positive lebesgue_integral_def diff_minus)
 qed
 
 lemma (in measure_space) integral_mono:
@@ -1795,20 +1905,21 @@
   by auto
 
 lemma (in measure_space) integral_indicator[simp, intro]:
-  assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
-  shows "integral\<^isup>L M (indicator a) = real (\<mu> a)" (is ?int)
-  and "integrable M (indicator a)" (is ?able)
+  assumes "A \<in> sets M" and "\<mu> A \<noteq> \<infinity>"
+  shows "integral\<^isup>L M (indicator A) = real (\<mu> A)" (is ?int)
+  and "integrable M (indicator A)" (is ?able)
 proof -
-  have *:
-    "\<And>A x. Real (indicator A x) = indicator A x"
-    "\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto
+  from `A \<in> sets M` have *:
+    "\<And>x. extreal (indicator A x) = indicator A x"
+    "(\<integral>\<^isup>+x. extreal (- indicator A x) \<partial>M) = 0"
+    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_extreal_def)
   show ?int ?able
     using assms unfolding lebesgue_integral_def integrable_def
     by (auto simp: * positive_integral_indicator borel_measurable_indicator)
 qed
 
 lemma (in measure_space) integral_cmul_indicator:
-  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>"
+  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<infinity>"
   shows "integrable M (\<lambda>x. c * indicator A x)" (is ?P)
   and "(\<integral>x. c * indicator A x \<partial>M) = c * real (\<mu> A)" (is ?I)
 proof -
@@ -1840,15 +1951,11 @@
   assumes "integrable M f"
   shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
 proof -
-  have *:
-    "\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)"
-    "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
-  have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and
-    f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
-        "(\<lambda>x. Real (- f x)) \<in> borel_measurable M"
-    using assms unfolding integrable_def by auto
-  from abs assms show ?thesis unfolding integrable_def *
-    using positive_integral_linear[OF f, of 1] by simp
+  from assms have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>)\<partial>M) = 0"
+    "\<And>x. extreal \<bar>f x\<bar> = max 0 (extreal (f x)) + max 0 (extreal (- f x))"
+    by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
+  with assms show ?thesis
+    by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
 qed
 
 lemma (in measure_space) integral_subalgebra:
@@ -1858,12 +1965,13 @@
     and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
 proof -
   interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
-  have "(\<lambda>x. Real (f x)) \<in> borel_measurable N" "(\<lambda>x. Real (- f x)) \<in> borel_measurable N"
-    using borel by auto
-  note * = this[THEN positive_integral_subalgebra[OF _ N sa]]
-  have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
+  have "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M)"
+       "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)"
+    using borel by (auto intro!: positive_integral_subalgebra N sa)
+  moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
     using assms unfolding measurable_def by auto
-  then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def)
+  ultimately show ?P ?I
+    by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0)
 qed
 
 lemma (in measure_space) integrable_bound:
@@ -1873,21 +1981,21 @@
   assumes borel: "g \<in> borel_measurable M"
   shows "integrable M g"
 proof -
-  have "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real \<bar>g x\<bar> \<partial>M)"
+  have "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal \<bar>g x\<bar> \<partial>M)"
     by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
     using f by (auto intro!: positive_integral_mono)
-  also have "\<dots> < \<omega>"
-    using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
-  finally have pos: "(\<integral>\<^isup>+ x. Real (g x) \<partial>M) < \<omega>" .
+  also have "\<dots> < \<infinity>"
+    using `integrable M f` unfolding integrable_def by auto
+  finally have pos: "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) < \<infinity>" .
 
-  have "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (\<bar>g x\<bar>) \<partial>M)"
+  have "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal (\<bar>g x\<bar>) \<partial>M)"
     by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. Real (f x) \<partial>M)"
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
     using f by (auto intro!: positive_integral_mono)
-  also have "\<dots> < \<omega>"
-    using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
-  finally have neg: "(\<integral>\<^isup>+ x. Real (- g x) \<partial>M) < \<omega>" .
+  also have "\<dots> < \<infinity>"
+    using `integrable M f` unfolding integrable_def by auto
+  finally have neg: "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) < \<infinity>" .
 
   from neg pos borel show ?thesis
     unfolding integrable_def by auto
@@ -1959,41 +2067,34 @@
       by (simp add: mono_def incseq_def) }
   note pos_u = this
 
-  hence [simp]: "\<And>i x. Real (- f i x) = 0" "\<And>x. Real (- u x) = 0"
-    using pos by auto
+  have SUP_F: "\<And>x. (SUP n. extreal (f n x)) = extreal (u x)"
+    unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
 
-  have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)"
-    using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2])
-
-  have borel_f: "\<And>i. (\<lambda>x. Real (f i x)) \<in> borel_measurable M"
+  have borel_f: "\<And>i. (\<lambda>x. extreal (f i x)) \<in> borel_measurable M"
     using i unfolding integrable_def by auto
-  hence "(\<lambda>x. SUP i. Real (f i x)) \<in> borel_measurable M"
+  hence "(\<lambda>x. SUP i. extreal (f i x)) \<in> borel_measurable M"
     by auto
   hence borel_u: "u \<in> borel_measurable M"
-    using pos_u by (auto simp: borel_measurable_Real_eq SUP_F)
+    by (auto simp: borel_measurable_extreal_iff SUP_F)
 
-  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M) = Real (integral\<^isup>L M (f n))"
-    using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real)
+  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. extreal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- u x) \<partial>M) = 0"
+    using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
+
+  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M) = extreal (integral\<^isup>L M (f n))"
+    using i positive_integral_positive by (auto simp: extreal_real lebesgue_integral_def integrable_def)
 
   have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
     using pos i by (auto simp: integral_positive)
   hence "0 \<le> x"
     using LIMSEQ_le_const[OF ilim, of 0] by auto
 
-  have "(\<lambda>i. (\<integral>\<^isup>+ x. Real (f i x) \<partial>M)) \<up> (\<integral>\<^isup>+ x. Real (u x) \<partial>M)"
-  proof (rule positive_integral_isoton)
-    from SUP_F mono pos
-    show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))"
-      unfolding isoton_fun_expand by (auto simp: isoton_def mono_def)
-  qed (rule borel_f)
-  hence pI: "(\<integral>\<^isup>+ x. Real (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. Real (f n x) \<partial>M))"
-    unfolding isoton_def by simp
-  also have "\<dots> = Real x" unfolding integral_eq
+  from mono pos i have pI: "(\<integral>\<^isup>+ x. extreal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M))"
+    by (auto intro!: positive_integral_monotone_convergence_SUP
+      simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
+  also have "\<dots> = extreal x" unfolding integral_eq
   proof (rule SUP_eq_LIMSEQ[THEN iffD2])
     show "mono (\<lambda>n. integral\<^isup>L M (f n))"
       using mono i by (auto simp: mono_def intro!: integral_mono)
-    show "\<And>n. 0 \<le> integral\<^isup>L M (f n)" using pos_integral .
-    show "0 \<le> x" using `0 \<le> x` .
     show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
   qed
   finally show  "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \<le> x`
@@ -2028,61 +2129,72 @@
   assumes "integrable M f"
   shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
 proof -
-  have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
+  have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>) \<partial>M) = 0"
+    using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
   have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
-  hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M"
-    "(\<integral>\<^isup>+ x. Real \<bar>f x\<bar> \<partial>M) \<noteq> \<omega>" unfolding integrable_def by auto
+  hence "(\<lambda>x. extreal (\<bar>f x\<bar>)) \<in> borel_measurable M"
+    "(\<integral>\<^isup>+ x. extreal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
   from positive_integral_0_iff[OF this(1)] this(2)
   show ?thesis unfolding lebesgue_integral_def *
-    by (simp add: real_of_pextreal_eq_0)
+    using positive_integral_positive[of "\<lambda>x. extreal \<bar>f x\<bar>"]
+    by (auto simp add: real_of_extreal_eq_0)
 qed
 
-lemma (in measure_space) positive_integral_omega:
-  assumes "f \<in> borel_measurable M"
-  and "integral\<^isup>P M f \<noteq> \<omega>"
-  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+lemma (in measure_space) positive_integral_PInf:
+  assumes f: "f \<in> borel_measurable M"
+  and not_Inf: "integral\<^isup>P M f \<noteq> \<infinity>"
+  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
 proof -
-  have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = (\<integral>\<^isup>+ x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x \<partial>M)"
-    using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
-  also have "\<dots> \<le> integral\<^isup>P M f"
-    by (auto intro!: positive_integral_mono simp: indicator_def)
-  finally show ?thesis
-    using assms(2) by (cases ?thesis) auto
+  have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+    using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets)
+  also have "\<dots> \<le> integral\<^isup>P M (\<lambda>x. max 0 (f x))"
+    by (auto intro!: positive_integral_mono simp: indicator_def max_def)
+  finally have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
+    by (simp add: positive_integral_max_0)
+  moreover have "0 \<le> \<mu> (f -` {\<infinity>} \<inter> space M)"
+    using f by (simp add: measurable_sets)
+  ultimately show ?thesis
+    using assms by (auto split: split_if_asm)
 qed
 
-lemma (in measure_space) positive_integral_omega_AE:
-  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>"
+lemma (in measure_space) positive_integral_PInf_AE:
+  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x. f x \<noteq> \<infinity>"
 proof (rule AE_I)
-  show "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
-    by (rule positive_integral_omega[OF assms])
-  show "f -` {\<omega>} \<inter> space M \<in> sets M"
+  show "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+    by (rule positive_integral_PInf[OF assms])
+  show "f -` {\<infinity>} \<inter> space M \<in> sets M"
     using assms by (auto intro: borel_measurable_vimage)
 qed auto
 
-lemma (in measure_space) simple_integral_omega:
-  assumes "simple_function M f"
-  and "integral\<^isup>S M f \<noteq> \<omega>"
-  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
-proof (rule positive_integral_omega)
+lemma (in measure_space) simple_integral_PInf:
+  assumes "simple_function M f" "\<And>x. 0 \<le> f x"
+  and "integral\<^isup>S M f \<noteq> \<infinity>"
+  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+proof (rule positive_integral_PInf)
   show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
-  show "integral\<^isup>P M f \<noteq> \<omega>"
+  show "integral\<^isup>P M f \<noteq> \<infinity>"
     using assms by (simp add: positive_integral_eq_simple_integral)
 qed
 
 lemma (in measure_space) integral_real:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  assumes [simp]: "AE x. f x \<noteq> \<omega>"
-  shows "(\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f)" (is ?plus)
-    and "(\<integral>x. - real (f x) \<partial>M) = - real (integral\<^isup>P M f)" (is ?minus)
-proof -
-  have "(\<integral>\<^isup>+ x. Real (real (f x)) \<partial>M) = integral\<^isup>P M f"
-    by (auto intro!: positive_integral_cong_AE simp: Real_real)
-  moreover
-  have "(\<integral>\<^isup>+ x. Real (- real (f x)) \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
-    by (intro positive_integral_cong) auto
-  ultimately show ?plus ?minus
-    by (auto simp: lebesgue_integral_def integrable_def)
-qed
+  "AE x. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
+  using assms unfolding lebesgue_integral_def
+  by (subst (1 2) positive_integral_cong_AE) (auto simp add: extreal_real)
+
+lemma liminf_extreal_cminus:
+  fixes f :: "nat \<Rightarrow> extreal" assumes "c \<noteq> -\<infinity>"
+  shows "liminf (\<lambda>x. c - f x) = c - limsup f"
+proof (cases c)
+  case PInf then show ?thesis by (simp add: Liminf_const)
+next
+  case (real r) then show ?thesis
+    unfolding liminf_SUPR_INFI limsup_INFI_SUPR
+    apply (subst INFI_extreal_cminus)
+    apply auto
+    apply (subst SUPR_extreal_cminus)
+    apply auto
+    done
+qed (insert `c \<noteq> -\<infinity>`, simp)
 
 lemma (in measure_space) integral_dominated_convergence:
   assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
@@ -2129,61 +2241,76 @@
     finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
   note diff_less_2w = this
 
-  have PI_diff: "\<And>m n. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M) =
-    (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
+  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. extreal (?diff n x) \<partial>M) =
+    (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
     using diff w diff_less_2w w_pos
     by (subst positive_integral_diff[symmetric])
        (auto simp: integrable_def intro!: positive_integral_cong)
 
   have "integrable M (\<lambda>x. 2 * w x)"
     using w by (auto intro: integral_cmult)
-  hence I2w_fin: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> \<omega>" and
-    borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M"
+  hence I2w_fin: "(\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
+    borel_2w: "(\<lambda>x. extreal (2 * w x)) \<in> borel_measurable M"
     unfolding integrable_def by auto
 
-  have "(INF n. SUP m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) = 0" (is "?lim_SUP = 0")
+  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
   proof cases
-    assume eq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = 0"
-    have "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) \<le> (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M)"
-    proof (rule positive_integral_mono)
-      fix i x assume "x \<in> space M" from diff_less_2w[OF this, of i]
-      show "Real \<bar>u i x - u' x\<bar> \<le> Real (2 * w x)" by auto
-    qed
-    hence "\<And>i. (\<integral>\<^isup>+ x. Real \<bar>u i x - u' x\<bar> \<partial>M) = 0" using eq_0 by auto
-    thus ?thesis by simp
+    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
+    { fix n
+      have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
+        using diff_less_2w[of _ n] unfolding positive_integral_max_0
+        by (intro positive_integral_mono) auto
+      then have "?f n = 0"
+        using positive_integral_positive[of ?f'] eq_0 by auto }
+    then show ?thesis by (simp add: Limsup_const)
   next
-    assume neq_0: "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) \<noteq> 0"
-    have "(\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \<partial>M)"
-    proof (rule positive_integral_cong, subst add_commute)
+    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
+    have "0 = limsup (\<lambda>n. 0 :: extreal)" by (simp add: Limsup_const)
+    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+      by (intro limsup_mono positive_integral_positive)
+    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)" .
+    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (extreal (?diff n x))) \<partial>M)"
+    proof (rule positive_integral_cong)
       fix x assume x: "x \<in> space M"
-      show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
-      proof (rule LIMSEQ_imp_lim_INF[symmetric])
-        fix j show "0 \<le> ?diff j x" using diff_less_2w[OF x, of j] by simp
-      next
+      show "max 0 (extreal (2 * w x)) = liminf (\<lambda>n. max 0 (extreal (?diff n x)))"
+        unfolding extreal_max_0
+      proof (rule lim_imp_Liminf[symmetric], unfold lim_extreal)
         have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
           using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
-        thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
-      qed
+        then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
+          by (auto intro!: tendsto_real_max simp add: lim_extreal)
+      qed (rule trivial_limit_sequentially)
     qed
-    also have "\<dots> \<le> (SUP n. INF m. (\<integral>\<^isup>+ x. Real (?diff (m + n) x) \<partial>M))"
+    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (extreal (?diff n x)) \<partial>M)"
       using u'_borel w u unfolding integrable_def
-      by (auto intro!: positive_integral_lim_INF)
-    also have "\<dots> = (\<integral>\<^isup>+ x. Real (2 * w x) \<partial>M) -
-        (INF n. SUP m. \<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)"
-      unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus ..
-    finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0)
+      by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
+    also have "\<dots> = (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) -
+        limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+      unfolding PI_diff positive_integral_max_0
+      using positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"]
+      by (subst liminf_extreal_cminus) auto
+    finally show ?thesis
+      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"] pos
+      unfolding positive_integral_max_0
+      by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"])
+         auto
   qed
 
-  have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto
-
-  have [simp]: "\<And>n m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M) =
-    Real ((\<integral>x. \<bar>u (n + m) x - u' x\<bar> \<partial>M))"
-    using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real)
-
-  have "(SUP n. INF m. (\<integral>\<^isup>+ x. Real \<bar>u (m + n) x - u' x\<bar> \<partial>M)) \<le> ?lim_SUP"
-    (is "?lim_INF \<le> _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP)
-  hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto
-  thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP)
+  have "liminf ?f \<le> limsup ?f"
+    by (intro extreal_Liminf_le_Limsup trivial_limit_sequentially)
+  moreover
+  { have "0 = liminf (\<lambda>n. 0 :: extreal)" by (simp add: Liminf_const)
+    also have "\<dots> \<le> liminf ?f"
+      by (intro liminf_mono positive_integral_positive)
+    finally have "0 \<le> liminf ?f" . }
+  ultimately have liminf_limsup_eq: "liminf ?f = extreal 0" "limsup ?f = extreal 0"
+    using `limsup ?f = 0` by auto
+  have "\<And>n. (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = extreal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
+    using diff positive_integral_positive
+    by (subst integral_eq_positive_integral) (auto simp: extreal_real integrable_def)
+  then show ?lim_diff
+    using extreal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
+    by (simp add: lim_extreal)
 
   show ?lim
   proof (rule LIMSEQ_I)
@@ -2266,7 +2393,7 @@
   assumes f: "f \<in> borel_measurable M"
   and bij: "bij_betw enum S (f ` space M)"
   and enum_zero: "enum ` (-S) \<subseteq> {0}"
-  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
   and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
   shows "integrable M f"
   and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
@@ -2303,7 +2430,7 @@
     also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
       using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
     finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
-      by (simp add: abs_mult_pos real_pextreal_pos) }
+      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_extreal_pos measurable_sets) }
   note int_abs_F = this
 
   have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
@@ -2323,7 +2450,7 @@
 
 lemma (in measure_space) integral_on_finite:
   assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
-  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
   shows "integrable M f"
   and "(\<integral>x. f x \<partial>M) =
     (\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral")
@@ -2353,11 +2480,12 @@
   by (auto intro: borel_measurable_simple_function)
 
 lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
-  "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
+  assumes pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+  shows "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
 proof -
   have *: "integral\<^isup>P M f = (\<integral>\<^isup>+ x. (\<Sum>y\<in>space M. f y * indicator {y} x) \<partial>M)"
     by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
-  show ?thesis unfolding * using borel_measurable_finite[of f]
+  show ?thesis unfolding * using borel_measurable_finite[of f] pos
     by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
 qed
 
@@ -2365,16 +2493,20 @@
   shows "integrable M f"
   and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
 proof -
-  have [simp]:
-    "(\<integral>\<^isup>+ x. Real (f x) \<partial>M) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
-    "(\<integral>\<^isup>+ x. Real (- f x) \<partial>M) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})"
-    unfolding positive_integral_finite_eq_setsum by auto
-  show "integrable M f" using finite_space finite_measure
-    by (simp add: setsum_\<omega> integrable_def)
-  show ?I using finite_measure
-    apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric]
-      real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric])
-    by (rule setsum_cong) (simp_all split: split_if)
+  have *:
+    "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (f x)) * \<mu> {x})"
+    "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (- f x)) * \<mu> {x})"
+    by (simp_all add: positive_integral_finite_eq_setsum)
+  then show "integrable M f" using finite_space finite_measure
+    by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
+             split: split_max)
+  show ?I using finite_measure *
+    apply (simp add: positive_integral_max_0 lebesgue_integral_def)
+    apply (subst (1 2) setsum_real_of_extreal[symmetric])
+    apply (simp_all split: split_max add: setsum_subtractf[symmetric])
+    apply (intro setsum_cong[OF refl])
+    apply (simp split: split_max)
+    done
 qed
 
 end
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -48,12 +48,12 @@
 lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
   unfolding Pi_def by auto
 
-subsection {* Lebesgue measure *}
+subsection {* Lebesgue measure *} 
 
 definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
   "lebesgue = \<lparr> space = UNIV,
     sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n},
-    measure = \<lambda>A. SUP n. Real (integral (cube n) (indicator A)) \<rparr>"
+    measure = \<lambda>A. SUP n. extreal (integral (cube n) (indicator A)) \<rparr>"
 
 lemma space_lebesgue[simp]: "space lebesgue = UNIV"
   unfolding lebesgue_def by simp
@@ -114,10 +114,33 @@
   qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
 qed simp
 
+lemma suminf_SUP_eq:
+  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> extreal"
+  assumes "\<And>i. incseq (\<lambda>n. f n i)" "\<And>n i. 0 \<le> f n i"
+  shows "(\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)"
+proof -
+  { fix n :: nat
+    have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
+      using assms by (auto intro!: SUPR_extreal_setsum[symmetric]) }
+  note * = this
+  show ?thesis using assms
+    apply (subst (1 2) suminf_extreal_eq_SUPR)
+    unfolding *
+    apply (auto intro!: le_SUPI2)
+    apply (subst SUP_commute) ..
+qed
+
 interpretation lebesgue: measure_space lebesgue
 proof
   have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
-  show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
+  show "positive lebesgue (measure lebesgue)"
+  proof (unfold positive_def, safe)
+    show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
+    fix A assume "A \<in> sets lebesgue"
+    then show "0 \<le> measure lebesgue A"
+      unfolding lebesgue_def
+      by (auto intro!: le_SUPI2 integral_nonneg)
+  qed
 next
   show "countably_additive lebesgue (measure lebesgue)"
   proof (intro countably_additive_def[THEN iffD2] allI impI)
@@ -130,23 +153,17 @@
     assume "(\<Union>i. A i) \<in> sets lebesgue"
     then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n"
       by (auto dest: lebesgueD)
-    show "(\<Sum>\<^isub>\<infinity>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
-    proof (simp add: lebesgue_def, subst psuminf_SUP_eq)
-      fix n i show "Real (?m n i) \<le> Real (?m (Suc n) i)"
-        using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le)
+    show "(\<Sum>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
+    proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI)
+      fix i n show "extreal (?m n i) \<le> extreal (?m (Suc n) i)"
+        using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
     next
-      show "(SUP n. \<Sum>\<^isub>\<infinity>i. Real (?m n i)) = (SUP n. Real (?M n UNIV))"
-        unfolding psuminf_def
-      proof (subst setsum_Real, (intro arg_cong[where f="SUPR UNIV"] ext ballI nn SUP_eq_LIMSEQ[THEN iffD2])+)
-        fix n :: nat show "mono (\<lambda>m. \<Sum>x<m. ?m n x)"
-        proof (intro mono_iff_le_Suc[THEN iffD2] allI)
-          fix m show "(\<Sum>x<m. ?m n x) \<le> (\<Sum>x<Suc m. ?m n x)"
-            using nn[of n m] by auto
-        qed
-        show "0 \<le> ?M n UNIV"
-          using UN_A by (auto intro!: integral_nonneg)
-        fix m show "0 \<le> (\<Sum>x<m. ?m n x)" by (auto intro!: setsum_nonneg)
-      next
+      fix i n show "0 \<le> extreal (?m n i)"
+        using rA unfolding lebesgue_def
+        by (auto intro!: le_SUPI2 integral_nonneg)
+    next
+      show "(SUP n. \<Sum>i. extreal (?m n i)) = (SUP n. extreal (?M n UNIV))"
+      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_extreal[THEN iffD2] sums_def[THEN iffD2])
         fix n
         have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
         from lebesgueD[OF this]
@@ -171,8 +188,8 @@
             ultimately show ?case
               using Suc A by (simp add: integral_add[symmetric])
           qed auto }
-        ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV"
-          by simp
+        ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"
+          by (simp add: atLeast0LessThan)
       qed
     qed
   qed
@@ -232,13 +249,11 @@
 
 lemma lmeasure_iff_LIMSEQ:
   assumes "A \<in> sets lebesgue" "0 \<le> m"
-  shows "lebesgue.\<mu> A = Real m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
+  shows "lebesgue.\<mu> A = extreal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
 proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ)
   show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
     using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
-  fix n show "0 \<le> integral (cube n) (indicator A::_=>real)"
-    using assms by (auto dest!: lebesgueD intro!: integral_nonneg)
-qed fact
+qed
 
 lemma has_integral_indicator_UNIV:
   fixes s A :: "'a::ordered_euclidean_space set" and x :: real
@@ -260,7 +275,7 @@
 
 lemma lmeasure_finite_has_integral:
   fixes s :: "'a::ordered_euclidean_space set"
-  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = Real m" "0 \<le> m"
+  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = extreal m" "0 \<le> m"
   shows "(indicator s has_integral m) UNIV"
 proof -
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
@@ -302,12 +317,14 @@
     unfolding m by (intro integrable_integral **)
 qed
 
-lemma lmeasure_finite_integrable: assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s \<noteq> \<omega>"
+lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "lebesgue.\<mu> s \<noteq> \<infinity>"
   shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
 proof (cases "lebesgue.\<mu> s")
-  case (preal m) from lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
+  case (real m)
+  with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
+    lebesgue.positive_measure[OF s]
   show ?thesis unfolding integrable_on_def by auto
-qed (insert assms, auto)
+qed (insert assms lebesgue.positive_measure[OF s], auto)
 
 lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
   shows "s \<in> sets lebesgue"
@@ -321,7 +338,7 @@
 qed
 
 lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
-  shows "lebesgue.\<mu> s = Real m"
+  shows "lebesgue.\<mu> s = extreal m"
 proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
   show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
@@ -346,28 +363,28 @@
 qed
 
 lemma has_integral_iff_lmeasure:
-  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m)"
+  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m)"
 proof
   assume "(indicator A has_integral m) UNIV"
   with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
-  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m"
+  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
     by (auto intro: has_integral_nonneg)
 next
-  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = Real m"
+  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
   then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
 qed
 
 lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
-  shows "lebesgue.\<mu> s = Real (integral UNIV (indicator s))"
+  shows "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))"
   using assms unfolding integrable_on_def
 proof safe
   fix y :: real assume "(indicator s has_integral y) UNIV"
   from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
-  show "lebesgue.\<mu> s = Real (integral UNIV (indicator s))" by simp
+  show "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))" by simp
 qed
 
 lemma lebesgue_simple_function_indicator:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
   assumes f:"simple_function lebesgue f"
   shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
   by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
@@ -376,7 +393,7 @@
   "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (lebesgue.\<mu> s)"
   by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
 
-lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<omega>"
+lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<infinity>"
   using lmeasure_eq_integral[OF assms] by auto
 
 lemma negligible_iff_lebesgue_null_sets:
@@ -409,37 +426,29 @@
   shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
   by (rule integral_unique) (rule has_integral_const)
 
-lemma lmeasure_UNIV[intro]: "lebesgue.\<mu> (UNIV::'a::ordered_euclidean_space set) = \<omega>"
-proof (simp add: lebesgue_def SUP_\<omega>, intro allI impI)
-  fix x assume "x < \<omega>"
-  then obtain r where r: "x = Real r" "0 \<le> r" by (cases x) auto
-  then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto
-  show "\<exists>i. x < Real (integral (cube i) (indicator UNIV::'a\<Rightarrow>real))"
-  proof (intro exI[of _ n])
-    have [simp]: "indicator UNIV = (\<lambda>x. 1)" by (auto simp: fun_eq_iff)
-    { fix m :: nat assume "0 < m" then have "real n \<le> (\<Prod>x<m. 2 * real n)"
-      proof (induct m)
-        case (Suc m)
-        show ?case
-        proof cases
-          assume "m = 0" then show ?thesis by (simp add: lessThan_Suc)
-        next
-          assume "m \<noteq> 0" then have "real n \<le> (\<Prod>x<m. 2 * real n)" using Suc by auto
-          then show ?thesis
-            by (auto simp: lessThan_Suc field_simps mult_le_cancel_left1)
-        qed
-      qed auto } note this[OF DIM_positive[where 'a='a], simp]
-    then have [simp]: "0 \<le> (\<Prod>x<DIM('a). 2 * real n)" using real_of_nat_ge_zero by arith
-    have "x < Real (of_nat n)" using n r by auto
-    also have "Real (of_nat n) \<le> Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
-      by (auto simp: real_eq_of_nat[symmetric] cube_def content_closed_interval_cases)
-    finally show "x < Real (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))" .
-  qed
-qed
+lemma lmeasure_UNIV[intro]: "lebesgue.\<mu> (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
+proof (simp add: lebesgue_def, intro SUP_PInfty bexI)
+  fix n :: nat
+  have "indicator UNIV = (\<lambda>x::'a. 1 :: real)" by auto
+  moreover
+  { have "real n \<le> (2 * real n) ^ DIM('a)"
+    proof (cases n)
+      case 0 then show ?thesis by auto
+    next
+      case (Suc n')
+      have "real n \<le> (2 * real n)^1" by auto
+      also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
+        using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
+      finally show ?thesis .
+    qed }
+  ultimately show "extreal (real n) \<le> extreal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
+    using integral_const DIM_positive[where 'a='a]
+    by (auto simp: cube_def content_closed_interval_cases setprod_constant)
+qed simp
 
 lemma
   fixes a b ::"'a::ordered_euclidean_space"
-  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = Real (content {a..b})"
+  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = extreal (content {a..b})"
 proof -
   have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
     unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw)
@@ -467,7 +476,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanAtMost[simp]:
-    "lebesgue.\<mu> {a <.. b} = Real (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a <.. b} = extreal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
@@ -479,7 +488,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_atLeastLessThan[simp]:
-    "lebesgue.\<mu> {a ..< b} = Real (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a ..< b} = extreal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
@@ -491,7 +500,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanLessThan[simp]:
-    "lebesgue.\<mu> {a <..< b} = Real (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a <..< b} = extreal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
@@ -511,19 +520,16 @@
   and measurable_lborel[simp]: "measurable lborel = measurable borel"
   by (simp_all add: measurable_def_raw lborel_def)
 
-interpretation lborel: measure_space lborel
+interpretation lborel: measure_space "lborel :: ('a::ordered_euclidean_space) measure_space"
   where "space lborel = UNIV"
   and "sets lborel = sets borel"
   and "measure lborel = lebesgue.\<mu>"
   and "measurable lborel = measurable borel"
-proof -
-  show "measure_space lborel"
-  proof
-    show "countably_additive lborel (measure lborel)"
-      using lebesgue.ca unfolding countably_additive_def lborel_def
-      apply safe apply (erule_tac x=A in allE) by auto
-  qed (auto simp: lborel_def)
-qed simp_all
+proof (rule lebesgue.measure_space_subalgebra)
+  have "sigma_algebra (lborel::'a measure_space) \<longleftrightarrow> sigma_algebra (borel::'a algebra)"
+    unfolding sigma_algebra_iff2 lborel_def by simp
+  then show "sigma_algebra (lborel::'a measure_space)" by simp default
+qed auto
 
 interpretation lborel: sigma_finite_measure lborel
   where "space lborel = UNIV"
@@ -536,7 +542,7 @@
     show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)
     { fix x have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
     thus "(\<Union>i. cube i) = space lborel" by auto
-    show "\<forall>i. measure lborel (cube i) \<noteq> \<omega>" by (simp add: cube_def)
+    show "\<forall>i. measure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
   qed
 qed simp_all
 
@@ -544,171 +550,221 @@
 proof
   from lborel.sigma_finite guess A ..
   moreover then have "range A \<subseteq> sets lebesgue" using lebesgueI_borel by blast
-  ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lebesgue.\<mu> (A i) \<noteq> \<omega>)"
+  ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lebesgue.\<mu> (A i) \<noteq> \<infinity>)"
     by auto
 qed
 
 subsection {* Lebesgue integrable implies Gauge integrable *}
 
+lemma positive_not_Inf:
+  "0 \<le> x \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> \<bar>x\<bar> \<noteq> \<infinity>"
+  by (cases x) auto
+
+lemma has_integral_cmult_real:
+  fixes c :: real
+  assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A"
+  shows "((\<lambda>x. c * f x) has_integral c * x) A"
+proof cases
+  assume "c \<noteq> 0"
+  from has_integral_cmul[OF assms[OF this], of c] show ?thesis
+    unfolding real_scaleR_def .
+qed simp
+
 lemma simple_function_has_integral:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
   assumes f:"simple_function lebesgue f"
-  and f':"\<forall>x. f x \<noteq> \<omega>"
-  and om:"\<forall>x\<in>range f. lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+  and f':"range f \<subseteq> {0..<\<infinity>}"
+  and om:"\<And>x. x \<in> range f \<Longrightarrow> lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
-  unfolding simple_integral_def
-  apply(subst lebesgue_simple_function_indicator[OF f])
-proof -
-  case goal1
-  have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
-    "\<forall>x\<in>range f. x * lebesgue.\<mu> (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
-    using f' om unfolding indicator_def by auto
-  show ?case unfolding space_lebesgue real_of_pextreal_setsum'[OF *(1),THEN sym]
-    unfolding real_of_pextreal_setsum'[OF *(2),THEN sym]
-    unfolding real_of_pextreal_setsum space_lebesgue
-    apply(rule has_integral_setsum)
-  proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
-    fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
-      real (f y * lebesgue.\<mu> (f -` {f y} \<inter> UNIV))) UNIV"
-    proof(cases "f y = 0") case False
-      have mea:"(indicator (f -` {f y}) ::_\<Rightarrow>real) integrable_on UNIV"
-        apply(rule lmeasure_finite_integrable)
-        using assms unfolding simple_function_def using False by auto
-      have *:"\<And>x. real (indicator (f -` {f y}) x::pextreal) = (indicator (f -` {f y}) x)"
-        by (auto simp: indicator_def)
-      show ?thesis unfolding real_of_pextreal_mult[THEN sym]
-        apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
-        unfolding Int_UNIV_right lmeasure_eq_integral[OF mea,THEN sym]
-        unfolding integral_eq_lmeasure[OF mea, symmetric] *
-        apply(rule integrable_integral) using mea .
-    qed auto
+  unfolding simple_integral_def space_lebesgue
+proof (subst lebesgue_simple_function_indicator)
+  let "?M x" = "lebesgue.\<mu> (f -` {x} \<inter> UNIV)"
+  let "?F x" = "indicator (f -` {x})"
+  { fix x y assume "y \<in> range f"
+    from subsetD[OF f' this] have "y * ?F y x = extreal (real y * ?F y x)"
+      by (cases rule: extreal2_cases[of y "?F y x"])
+         (auto simp: indicator_def one_extreal_def split: split_if_asm) }
+  moreover
+  { fix x assume x: "x\<in>range f"
+    have "x * ?M x = real x * real (?M x)"
+    proof cases
+      assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
+      with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis
+        by (cases rule: extreal2_cases[of x "?M x"]) auto
+    qed simp }
+  ultimately
+  have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow>
+    ((\<lambda>x. \<Sum>y\<in>range f. real y * ?F y x) has_integral (\<Sum>x\<in>range f. real x * real (?M x))) UNIV"
+    by simp
+  also have \<dots>
+  proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
+               real_of_extreal_pos lebesgue.positive_measure ballI)
+    show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue" "\<And>y. f -` {y} \<inter> UNIV \<in> sets lebesgue"
+      using lebesgue.simple_functionD[OF f] by auto
+    fix y assume "real y \<noteq> 0" "y \<in> range f"
+    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = extreal (real (?M y))"
+      by (auto simp: extreal_real)
   qed
-qed
+  finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .
+qed fact
 
 lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
   unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
   using assms by auto
 
 lemma simple_function_has_integral':
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
-  assumes f:"simple_function lebesgue f"
-  and i: "integral\<^isup>S lebesgue f \<noteq> \<omega>"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
+  assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"
+  and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
-proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
-  { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
-  have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
-  have **:"lebesgue.\<mu> {x\<in>space lebesgue. f x \<noteq> ?f x} = 0"
-    using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
-  show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
-    apply(rule lebesgue.simple_function_compose1[OF f])
-    unfolding * defer apply(rule simple_function_has_integral)
-  proof-
-    show "simple_function lebesgue ?f"
-      using lebesgue.simple_function_compose1[OF f] .
-    show "\<forall>x. ?f x \<noteq> \<omega>" by auto
-    show "\<forall>x\<in>range ?f. lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
-    proof (safe, simp, safe, rule ccontr)
-      fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
-      hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
-        by (auto split: split_if_asm)
-      moreover assume "lebesgue.\<mu> ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
-      ultimately have "lebesgue.\<mu> (f -` {f y}) = \<omega>" by simp
-      moreover
-      have "f y * lebesgue.\<mu> (f -` {f y}) \<noteq> \<omega>" using i f
-        unfolding simple_integral_def setsum_\<omega> simple_function_def
-        by auto
-      ultimately have "f y = 0" by (auto split: split_if_asm)
-      then show False using `f y \<noteq> 0` by simp
-    qed
+proof -
+  let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
+  note f(1)[THEN lebesgue.simple_functionD(2)]
+  then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto
+  have f': "simple_function lebesgue ?f"
+    using f by (intro lebesgue.simple_function_If_set) auto
+  have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
+  have "AE x in lebesgue. f x = ?f x"
+    using lebesgue.simple_integral_PInf[OF f i]
+    by (intro lebesgue.AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
+  from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"
+    by (rule lebesgue.simple_integral_cong_AE)
+  have real_eq: "\<And>x. real (f x) = real (?f x)" by auto
+
+  show ?thesis
+    unfolding eq real_eq
+  proof (rule simple_function_has_integral[OF f' rng])
+    fix x assume x: "x \<in> range ?f" and inf: "lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = \<infinity>"
+    have "x * lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
+      using f'[THEN lebesgue.simple_functionD(2)]
+      by (simp add: lebesgue.simple_integral_cmult_indicator)
+    also have "\<dots> \<le> integral\<^isup>S lebesgue f"
+      using f'[THEN lebesgue.simple_functionD(2)] f
+      by (intro lebesgue.simple_integral_mono lebesgue.simple_function_mult lebesgue.simple_function_indicator)
+         (auto split: split_indicator)
+    finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm)
   qed
 qed
 
-lemma (in measure_space) positive_integral_monotone_convergence:
-  fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pextreal"
-  assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
-  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
-  shows "u \<in> borel_measurable M"
-  and "(\<lambda>i. integral\<^isup>P M (f i)) ----> integral\<^isup>P M u" (is ?ilim)
-proof -
-  from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
-  show ?ilim using mono lim i by auto
-  have "\<And>x. (SUP i. f i x) = u x" using mono lim SUP_Lim_pextreal
-    unfolding fun_eq_iff mono_def by auto
-  moreover have "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M"
-    using i by auto
-  ultimately show "u \<in> borel_measurable M" by simp
-qed
+lemma real_of_extreal_positive_mono:
+  "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
+  by (cases rule: extreal2_cases[of x y]) auto
 
 lemma positive_integral_has_integral:
-  fixes f::"'a::ordered_euclidean_space => pextreal"
-  assumes f:"f \<in> borel_measurable lebesgue"
-  and int_om:"integral\<^isup>P lebesgue f \<noteq> \<omega>"
-  and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
+  assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
-proof- let ?i = "integral\<^isup>P lebesgue f"
-  from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
-  guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
-  let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
-  have u_simple:"\<And>k. integral\<^isup>S lebesgue (u k) = integral\<^isup>P lebesgue (u k)"
-    apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
-  have int_u_le:"\<And>k. integral\<^isup>S lebesgue (u k) \<le> integral\<^isup>P lebesgue f"
-    unfolding u_simple apply(rule lebesgue.positive_integral_mono)
-    using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
-  have u_int_om:"\<And>i. integral\<^isup>S lebesgue (u i) \<noteq> \<omega>"
-  proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
+proof -
+  from lebesgue.borel_measurable_implies_simple_function_sequence'[OF f(1)]
+  guess u . note u = this
+  have SUP_eq: "\<And>x. (SUP i. u i x) = f x"
+    using u(4) f(2)[THEN subsetD] by (auto split: split_max)
+  let "?u i x" = "real (u i x)"
+  note u_eq = lebesgue.positive_integral_eq_simple_integral[OF u(1,5), symmetric]
+  { fix i
+    note u_eq
+    also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"
+      by (intro lebesgue.positive_integral_mono) (auto intro: le_SUPI simp: u(4)[symmetric])
+    finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"
+      unfolding positive_integral_max_0 using f by auto }
+  note u_fin = this
+  then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
+    by (rule simple_function_has_integral'[OF u(1,5)])
+  have "\<forall>x. \<exists>r\<ge>0. f x = extreal r"
+  proof
+    fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)
+    then show "\<exists>r\<ge>0. f x = extreal r" by (cases "f x") auto
+  qed
+  from choice[OF this] obtain f' where f': "f = (\<lambda>x. extreal (f' x))" "\<And>x. 0 \<le> f' x" by auto
+
+  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+  proof
+    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+    proof (intro choice allI)
+      fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis
+      then show "\<exists>r\<ge>0. u i x = extreal r" using u(5)[of i x] by (cases "u i x") auto
+    qed
+  qed
+  from choice[OF this] obtain u' where
+      u': "u = (\<lambda>i x. extreal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
 
-  note u_int = simple_function_has_integral'[OF u(1) this]
-  have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
-    (\<lambda>k. Integration.integral UNIV (\<lambda>x. real (u k x))) ----> Integration.integral UNIV (\<lambda>x. real (f x))"
-    apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
-  proof safe case goal1 show ?case apply(rule real_of_pextreal_mono) using u(2,3) by auto
-  next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
-      prefer 3 apply(subst Real_real') defer apply(subst Real_real')
-      using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
-  next case goal3
-    show ?case apply(rule bounded_realI[where B="real (integral\<^isup>P lebesgue f)"])
-      apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
-      unfolding integral_unique[OF u_int] defer apply(rule real_of_pextreal_mono[OF _ int_u_le])
-      using u int_om by auto
-  qed note int = conjunctD2[OF this]
+  have convergent: "f' integrable_on UNIV \<and>
+    (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"
+  proof (intro monotone_convergence_increasing allI ballI)
+    show int: "\<And>k. (u' k) integrable_on UNIV"
+      using u_int unfolding integrable_on_def u' by auto
+    show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)
+      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_extreal_positive_mono)
+    show "\<And>x. (\<lambda>k. u' k x) ----> f' x"
+      using SUP_eq u(2)
+      by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)
+    show "bounded {integral UNIV (u' k)|k. True}"
+    proof (safe intro!: bounded_realI)
+      fix k
+      have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)"
+        by (intro abs_of_nonneg integral_nonneg int ballI u')
+      also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"
+        using u_int[THEN integral_unique] by (simp add: u')
+      also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
+        using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp
+      also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
+        by (auto intro!: real_of_extreal_positive_mono lebesgue.positive_integral_positive
+             lebesgue.positive_integral_mono le_SUPI simp: SUP_eq[symmetric])
+      finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
+    qed
+  qed
 
-  have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> ?i" unfolding u_simple
-    apply(rule lebesgue.positive_integral_monotone_convergence(2))
-    apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
-    using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
-  hence "(\<lambda>i. real (integral\<^isup>S lebesgue (u i))) ----> real ?i" apply-
-    apply(subst lim_Real[THEN sym]) prefer 3
-    apply(subst Real_real') defer apply(subst Real_real')
-    using u f_om int_om u_int_om by auto
-  note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
-  show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
+  have "integral\<^isup>P lebesgue f = extreal (integral UNIV f')"
+  proof (rule tendsto_unique[OF trivial_limit_sequentially])
+    have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
+      unfolding u_eq by (intro LIMSEQ_extreal_SUPR lebesgue.incseq_positive_integral u)
+    also note lebesgue.positive_integral_monotone_convergence_SUP
+      [OF u(2)  lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric]
+    finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
+      unfolding SUP_eq .
+
+    { fix k
+      have "0 \<le> integral\<^isup>S lebesgue (u k)"
+        using u by (auto intro!: lebesgue.simple_integral_positive)
+      then have "integral\<^isup>S lebesgue (u k) = extreal (real (integral\<^isup>S lebesgue (u k)))"
+        using u_fin by (auto simp: extreal_real) }
+    note * = this
+    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> extreal (integral UNIV f')"
+      using convergent using u_int[THEN integral_unique, symmetric]
+      by (subst *) (simp add: lim_extreal u')
+  qed
+  then show ?thesis using convergent by (simp add: f' integrable_integral)
 qed
 
 lemma lebesgue_integral_has_integral:
-  fixes f::"'a::ordered_euclidean_space => real"
-  assumes f:"integrable lebesgue f"
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+  assumes f: "integrable lebesgue f"
   shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
-proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
-  have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
-  note f = integrableD[OF f]
-  show ?thesis unfolding lebesgue_integral_def apply(subst *)
-  proof(rule has_integral_sub) case goal1
-    have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
-    note lebesgue.borel_measurable_Real[OF f(1)]
-    from positive_integral_has_integral[OF this f(2) *]
-    show ?case unfolding real_Real_max .
-  next case goal2
-    have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
-    note lebesgue.borel_measurable_uminus[OF f(1)]
-    note lebesgue.borel_measurable_Real[OF this]
-    from positive_integral_has_integral[OF this f(3) *]
-    show ?case unfolding real_Real_max minus_min_eq_max by auto
-  qed
+proof -
+  let ?n = "\<lambda>x. real (extreal (max 0 (- f x)))" and ?p = "\<lambda>x. real (extreal (max 0 (f x)))"
+  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: extreal_max)
+  { fix f have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. extreal (max 0 (f x)) \<partial>lebesgue)"
+      by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) }
+  note eq = this
+  show ?thesis
+    unfolding lebesgue_integral_def
+    apply (subst *)
+    apply (rule has_integral_sub)
+    unfolding eq[of f] eq[of "\<lambda>x. - f x"]
+    apply (safe intro!: positive_integral_has_integral)
+    using integrableD[OF f]
+    by (auto simp: zero_extreal_def[symmetric] positive_integral_max_0  split: split_max
+             intro!: lebesgue.measurable_If lebesgue.borel_measurable_extreal)
 qed
 
 lemma lebesgue_positive_integral_eq_borel:
-  "f \<in> borel_measurable borel \<Longrightarrow> integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
-  by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
+  assumes f: "f \<in> borel_measurable borel"
+  shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
+proof -
+  from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"
+    by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
+  then show ?thesis unfolding positive_integral_max_0 .
+qed
 
 lemma lebesgue_integral_eq_borel:
   assumes "f \<in> borel_measurable borel"
@@ -771,7 +827,7 @@
   have "sets ?G = sets (\<Pi>\<^isub>M i\<in>I.
        sigma \<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>)"
     by (subst sigma_product_algebra_sigma_eq[of I "\<lambda>_ i. {..<real i}" ])
-       (auto intro!: measurable_sigma_sigma isotoneI real_arch_lt
+       (auto intro!: measurable_sigma_sigma incseq_SucI real_arch_lt
              simp: product_algebra_def)
   then show ?thesis
     unfolding lborel_def borel_eq_lessThan lebesgue_def sigma_def by simp
@@ -838,9 +894,10 @@
   let ?E = "\<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
   show "Int_stable ?E" using Int_stable_cuboids .
   show "range cube \<subseteq> sets ?E" unfolding cube_def_raw by auto
+  show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI)
   { fix x have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastsimp }
-  then show "cube \<up> space ?E" by (intro isotoneI cube_subset_Suc) auto
-  { fix i show "lborel.\<mu> (cube i) \<noteq> \<omega>" unfolding cube_def by auto }
+  then show "(\<Union>i. cube i) = space ?E" by auto
+  { fix i show "lborel.\<mu> (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
   show "A \<in> sets (sigma ?E)" "sets (sigma ?E) = sets lborel" "space ?E = space lborel"
     using assms by (simp_all add: borel_eq_atLeastAtMost)
 
@@ -857,7 +914,7 @@
         by (simp add: interval_ne_empty eucl_le[where 'a='a])
       then have "lborel.\<mu> X = (\<Prod>x<DIM('a). lborel.\<mu> {a $$ x .. b $$ x})"
         by (auto simp: content_closed_interval eucl_le[where 'a='a]
-                 intro!: Real_setprod )
+                 intro!: setprod_extreal[symmetric])
       also have "\<dots> = measure ?P (?T X)"
         unfolding * by (subst lborel_space.measure_times) auto
       finally show ?thesis .
@@ -882,7 +939,7 @@
   using lborel_eq_lborel_space[OF A] by simp
 
 lemma borel_fubini_positiv_integral:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pextreal"
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
   assumes f: "f \<in> borel_measurable borel"
   shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
 proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _])
--- a/src/HOL/Probability/Measure.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Measure.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -76,7 +76,7 @@
 
 lemma (in measure_space) measure_countably_additive:
   assumes "range A \<subseteq> sets M" "disjoint_family A"
-  shows "psuminf (\<lambda>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
+  shows "(\<Sum>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
 proof -
   have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN)
   with ca assms show ?thesis by (simp add: countably_additive_def)
@@ -94,13 +94,13 @@
   interpret N: sigma_algebra N by (intro sigma_algebra_cong assms)
   show ?thesis
   proof
-    show "measure N {} = 0" using assms by auto
+    show "positive N (measure N)" using assms by (auto simp: positive_def)
     show "countably_additive N (measure N)" unfolding countably_additive_def
     proof safe
       fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets N" "disjoint_family A"
       then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" unfolding assms by auto
       from measure_countably_additive[of A] A this[THEN assms(1)]
-      show "(\<Sum>\<^isub>\<infinity>n. measure N (A n)) = measure N (UNION UNIV A)"
+      show "(\<Sum>n. measure N (A n)) = measure N (UNION UNIV A)"
         unfolding assms by simp
     qed
   qed
@@ -124,51 +124,51 @@
   have "b = a \<union> (b - a)" using assms by auto
   moreover have "{} = a \<inter> (b - a)" by auto
   ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)"
-    using measure_additive[of a "b - a"] local.Diff[of b a] assms by auto
-  moreover have "\<mu> (b - a) \<ge> 0" using assms by auto
+    using measure_additive[of a "b - a"] Diff[of b a] assms by auto
+  moreover have "\<mu> a + 0 \<le> \<mu> a + \<mu> (b - a)" using assms by (intro add_mono) auto
   ultimately show "\<mu> a \<le> \<mu> b" by auto
 qed
 
 lemma (in measure_space) measure_compl:
-  assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<omega>"
+  assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<infinity>"
   shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
 proof -
   have s_less_space: "\<mu> s \<le> \<mu> (space M)"
     using s by (auto intro!: measure_mono sets_into_space)
-
+  from s have "0 \<le> \<mu> s" by auto
   have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s
     by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
   also have "... = \<mu> s + \<mu> (space M - s)"
     by (rule additiveD [OF additive]) (auto simp add: s)
   finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
-  thus ?thesis
-    unfolding minus_pextreal_eq2[OF s_less_space fin]
-    by (simp add: ac_simps)
+  then show ?thesis
+    using fin `0 \<le> \<mu> s`
+    unfolding extreal_eq_minus_iff by (auto simp: ac_simps)
 qed
 
 lemma (in measure_space) measure_Diff:
-  assumes finite: "\<mu> B \<noteq> \<omega>"
+  assumes finite: "\<mu> B \<noteq> \<infinity>"
   and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
   shows "\<mu> (A - B) = \<mu> A - \<mu> B"
 proof -
-  have *: "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
-
-  have "\<mu> ((A - B) \<union> B) = \<mu> (A - B) + \<mu> B"
-    using measurable by (rule_tac measure_additive[symmetric]) auto
-  thus ?thesis unfolding * using `\<mu> B \<noteq> \<omega>`
-    by (simp add: pextreal_cancel_plus_minus)
+  have "0 \<le> \<mu> B" using assms by auto
+  have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
+  then have "\<mu> A = \<mu> ((A - B) \<union> B)" by simp
+  also have "\<dots> = \<mu> (A - B) + \<mu> B"
+    using measurable by (subst measure_additive[symmetric]) auto
+  finally show "\<mu> (A - B) = \<mu> A - \<mu> B"
+    unfolding extreal_eq_minus_iff
+    using finite `0 \<le> \<mu> B` by auto
 qed
 
 lemma (in measure_space) measure_countable_increasing:
   assumes A: "range A \<subseteq> sets M"
       and A0: "A 0 = {}"
-      and ASuc: "\<And>n.  A n \<subseteq> A (Suc n)"
+      and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
   shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
 proof -
-  {
-    fix n
-    have "\<mu> (A n) =
-          setsum (\<mu> \<circ> (\<lambda>i. A (Suc i) - A i)) {..<n}"
+  { fix n
+    have "\<mu> (A n) = (\<Sum>i<n. \<mu> (A (Suc i) - A i))"
       proof (induct n)
         case 0 thus ?case by (auto simp add: A0)
       next
@@ -199,92 +199,83 @@
     by (metis A Diff range_subsetD)
   have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
     by (blast intro: range_subsetD [OF A])
-  have "psuminf ( (\<lambda>i. \<mu> (A (Suc i) - A i))) = \<mu> (\<Union>i. A (Suc i) - A i)"
+  have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))"
+    using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric])
+  also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)"
     by (rule measure_countably_additive)
        (auto simp add: disjoint_family_Suc ASuc A1 A2)
   also have "... =  \<mu> (\<Union>i. A i)"
     by (simp add: Aeq)
-  finally have "psuminf (\<lambda>i. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
-  thus ?thesis
-    by (auto simp add: Meq psuminf_def)
+  finally have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
+  then show ?thesis by (auto simp add: Meq)
 qed
 
 lemma (in measure_space) continuity_from_below:
-  assumes A: "range A \<subseteq> sets M"
-      and ASuc: "!!n.  A n \<subseteq> A (Suc n)"
+  assumes A: "range A \<subseteq> sets M" and "incseq A"
   shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
 proof -
   have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
-    apply (rule Sup_mono_offset_Suc)
-    apply (rule measure_mono)
-    using assms by (auto split: nat.split)
-
+    using A by (auto intro!: SUPR_eq exI split: nat.split)
   have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
     by (auto simp add: split: nat.splits)
   have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
     by simp
   have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
-    by (rule measure_countable_increasing)
-       (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits)
+    using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
+    by (force split: nat.splits intro!: measure_countable_increasing)
   also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
     by (simp add: ueq)
   finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
   thus ?thesis unfolding meq * comp_def .
 qed
 
-lemma (in measure_space) measure_up:
-  assumes "\<And>i. B i \<in> sets M" "B \<up> P"
-  shows "(\<lambda>i. \<mu> (B i)) \<up> \<mu> P"
-  using assms unfolding isoton_def
-  by (auto intro!: measure_mono continuity_from_below)
+lemma (in measure_space) measure_incseq:
+  assumes "range B \<subseteq> sets M" "incseq B"
+  shows "incseq (\<lambda>i. \<mu> (B i))"
+  using assms by (auto simp: incseq_def intro!: measure_mono)
 
-lemma (in measure_space) continuity_from_below':
-  assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
-  shows "(\<lambda>i. (\<mu> (A i))) ----> (\<mu> (\<Union>i. A i))"
-proof- let ?A = "\<Union>i. A i"
-  have " (\<lambda>i. \<mu> (A i)) \<up> \<mu> ?A" apply(rule measure_up)
-    using assms unfolding complete_lattice_class.isoton_def by auto
-  thus ?thesis by(rule isotone_Lim(1))
-qed
+lemma (in measure_space) continuity_from_below_Lim:
+  assumes A: "range A \<subseteq> sets M" "incseq A"
+  shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)"
+  using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A]
+    continuity_from_below[OF A] by simp
+
+lemma (in measure_space) measure_decseq:
+  assumes "range B \<subseteq> sets M" "decseq B"
+  shows "decseq (\<lambda>i. \<mu> (B i))"
+  using assms by (auto simp: decseq_def intro!: measure_mono)
 
 lemma (in measure_space) continuity_from_above:
-  assumes A: "range A \<subseteq> sets M"
-  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
-  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+  assumes A: "range A \<subseteq> sets M" and "decseq A"
+  and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
   shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)"
 proof -
-  { fix n have "A n \<subseteq> A 0" using mono_Suc by (induct n) auto }
-  note mono = this
-
   have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
     using A by (auto intro!: measure_mono)
-  hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using finite[of 0] by auto
+  hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
+
+  have A0: "0 \<le> \<mu> (A 0)" using A by auto
 
-  have le_IM: "(INF n. \<mu> (A n)) \<le> \<mu> (A 0)"
-    by (rule INF_leI) simp
-
-  have "\<mu> (A 0) - (INF n. \<mu> (A n)) = (SUP n. \<mu> (A 0 - A n))"
-    unfolding pextreal_SUP_minus[symmetric]
-    using mono A finite by (subst measure_Diff) auto
+  have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))"
+    by (simp add: extreal_SUPR_uminus minus_extreal_def)
+  also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))"
+    unfolding minus_extreal_def using A0 assms
+    by (subst SUPR_extreal_add) (auto simp add: measure_decseq)
+  also have "\<dots> = (SUP n. \<mu> (A 0 - A n))"
+    using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto
   also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
   proof (rule continuity_from_below)
     show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
       using A by auto
-    show "\<And>n. A 0 - A n \<subseteq> A 0 - A (Suc n)"
-      using mono_Suc by auto
+    show "incseq (\<lambda>n. A 0 - A n)"
+      using `decseq A` by (auto simp add: incseq_def decseq_def)
   qed
   also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
-    using mono A finite * by (simp, subst measure_Diff) auto
+    using A finite * by (simp, subst measure_Diff) auto
   finally show ?thesis
-    by (rule pextreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
+    unfolding extreal_minus_eq_minus_iff using finite A0 by auto
 qed
 
-lemma (in measure_space) measure_down:
-  assumes "\<And>i. B i \<in> sets M" "B \<down> P"
-  and finite: "\<And>i. \<mu> (B i) \<noteq> \<omega>"
-  shows "(\<lambda>i. \<mu> (B i)) \<down> \<mu> P"
-  using assms unfolding antiton_def
-  by (auto intro!: measure_mono continuity_from_above)
 lemma (in measure_space) measure_insert:
   assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
   shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
@@ -293,109 +284,26 @@
   from measure_additive[OF sets this] show ?thesis by simp
 qed
 
-lemma (in measure_space) measure_finite_singleton:
-  assumes fin: "finite S"
-  and ssets: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
+lemma (in measure_space) measure_setsum:
+  assumes "finite S" and "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
+  assumes disj: "disjoint_family_on A S"
+  shows "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>i\<in>S. A i)"
 using assms proof induct
-  case (insert x S)
-  have *: "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})" "{x} \<in> sets M"
-    using insert.prems by (blast intro!: insert.hyps(3))+
-
-  have "(\<Union>x\<in>S. {x}) \<in> sets M"
-    using  insert.prems `finite S` by (blast intro!: finite_UN)
-  hence "S \<in> sets M" by auto
-  from measure_insert[OF `{x} \<in> sets M` this `x \<notin> S`]
-  show ?case using `x \<notin> S` `finite S` * by simp
+  case (insert i S)
+  then have "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>a\<in>S. A a)"
+    by (auto intro: disjoint_family_on_mono)
+  moreover have "A i \<inter> (\<Union>a\<in>S. A a) = {}"
+    using `disjoint_family_on A (insert i S)` `i \<notin> S`
+    by (auto simp: disjoint_family_on_def)
+  ultimately show ?case using insert
+    by (auto simp: measure_additive finite_UN)
 qed simp
 
-lemma (in measure_space) measure_finitely_additive':
-  assumes "f \<in> ({..< n :: nat} \<rightarrow> sets M)"
-  assumes "\<And> a b. \<lbrakk>a < n ; b < n ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
-  assumes "s = \<Union> (f ` {..< n})"
-  shows "(\<Sum>i<n. (\<mu> \<circ> f) i) = \<mu> s"
-proof -
-  def f' == "\<lambda> i. (if i < n then f i else {})"
-  have rf: "range f' \<subseteq> sets M" unfolding f'_def
-    using assms empty_sets by auto
-  have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def
-    using assms by simp
-  have "\<Union> range f' = (\<Union> i \<in> {..< n}. f i)"
-    unfolding f'_def by auto
-  then have "\<mu> s = \<mu> (\<Union> range f')"
-    using assms by simp
-  then have part1: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = \<mu> s"
-    using df rf ca[unfolded countably_additive_def, rule_format, of f']
-    by auto
-
-  have "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum> i< n. \<mu> (f' i))"
-    by (rule psuminf_finite) (simp add: f'_def)
-  also have "\<dots> = (\<Sum>i<n. \<mu> (f i))"
-    unfolding f'_def by auto
-  finally have part2: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum>i<n. \<mu> (f i))" by simp
-  show ?thesis using part1 part2 by auto
-qed
-
-
-lemma (in measure_space) measure_finitely_additive:
-  assumes "finite r"
-  assumes "r \<subseteq> sets M"
-  assumes d: "\<And> a b. \<lbrakk>a \<in> r ; b \<in> r ; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<inter> b = {}"
-  shows "(\<Sum> i \<in> r. \<mu> i) = \<mu> (\<Union> r)"
-using assms
-proof -
-  (* counting the elements in r is enough *)
-  let ?R = "{..<card r}"
-  obtain f where f: "f ` ?R = r" "inj_on f ?R"
-    using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite r`]
-    unfolding atLeast0LessThan by auto
-  hence f_into_sets: "f \<in> ?R \<rightarrow> sets M" using assms by auto
-  have disj: "\<And> a b. \<lbrakk>a \<in> ?R ; b \<in> ?R ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
-  proof -
-    fix a b assume asm: "a \<in> ?R" "b \<in> ?R" "a \<noteq> b"
-    hence neq: "f a \<noteq> f b" using f[unfolded inj_on_def, rule_format] by blast
-    from asm have "f a \<in> r" "f b \<in> r" using f by auto
-    thus "f a \<inter> f b = {}" using d[of "f a" "f b"] f using neq by auto
-  qed
-  have "(\<Union> r) = (\<Union> i \<in> ?R. f i)"
-    using f by auto
-  hence "\<mu> (\<Union> r)= \<mu> (\<Union> i \<in> ?R. f i)" by simp
-  also have "\<dots> = (\<Sum> i \<in> ?R. \<mu> (f i))"
-    using measure_finitely_additive'[OF f_into_sets disj] by simp
-  also have "\<dots> = (\<Sum> a \<in> r. \<mu> a)"
-    using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. \<mu> a"] by auto
-  finally show ?thesis by simp
-qed
-
-lemma (in measure_space) measure_finitely_additive'':
-  assumes "finite s"
-  assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<in> sets M"
-  assumes d: "disjoint_family_on a s"
-  shows "(\<Sum> i \<in> s. \<mu> (a i)) = \<mu> (\<Union> i \<in> s. a i)"
-using assms
-proof -
-  (* counting the elements in s is enough *)
-  let ?R = "{..< card s}"
-  obtain f where f: "f ` ?R = s" "inj_on f ?R"
-    using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`]
-    unfolding atLeast0LessThan by auto
-  hence f_into_sets: "a \<circ> f \<in> ?R \<rightarrow> sets M" using assms unfolding o_def by auto
-  have disj: "\<And> i j. \<lbrakk>i \<in> ?R ; j \<in> ?R ; i \<noteq> j\<rbrakk> \<Longrightarrow> (a \<circ> f) i \<inter> (a \<circ> f) j = {}"
-  proof -
-    fix i j assume asm: "i \<in> ?R" "j \<in> ?R" "i \<noteq> j"
-    hence neq: "f i \<noteq> f j" using f[unfolded inj_on_def, rule_format] by blast
-    from asm have "f i \<in> s" "f j \<in> s" using f by auto
-    thus "(a \<circ> f) i \<inter> (a \<circ> f) j = {}"
-      using d f neq unfolding disjoint_family_on_def by auto
-  qed
-  have "(\<Union> i \<in> s. a i) = (\<Union> i \<in> f ` ?R. a i)" using f by auto
-  hence "(\<Union> i \<in> s. a i) = (\<Union> i \<in> ?R. a (f i))" by auto
-  hence "\<mu> (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. \<mu> (a (f i)))"
-    using measure_finitely_additive'[OF f_into_sets disj] by simp
-  also have "\<dots> = (\<Sum> i \<in> s. \<mu> (a i))"
-    using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. \<mu> (a i)"] by auto
-  finally show ?thesis by simp
-qed
+lemma (in measure_space) measure_finite_singleton:
+  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
+  using measure_setsum[of S "\<lambda>x. {x}", OF assms]
+  by (auto simp: disjoint_family_on_def)
 
 lemma finite_additivity_sufficient:
   assumes "sigma_algebra M"
@@ -405,7 +313,7 @@
   interpret sigma_algebra M by fact
   show ?thesis
   proof
-    show [simp]: "measure M {} = 0" using pos by (simp add: positive_def)
+    show [simp]: "positive M (measure M)" using pos by (simp add: positive_def)
     show "countably_additive M (measure M)"
     proof (auto simp add: countably_additive_def)
       fix A :: "nat \<Rightarrow> 'a set"
@@ -434,12 +342,12 @@
             by blast
         qed
       then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
-      then have "\<forall>m\<ge>N. measure M (A m) = 0" by simp
-      then have "(\<Sum>\<^isub>\<infinity> n. measure M (A n)) = setsum (\<lambda>m. measure M (A m)) {..<N}"
-        by (simp add: psuminf_finite)
+      then have "\<forall>m\<ge>N. measure M (A m) = 0" using pos[unfolded positive_def] by simp
+      then have "(\<Sum>n. measure M (A n)) = (\<Sum>m<N. measure M (A m))"
+        by (simp add: suminf_finite)
       also have "... = measure M (\<Union>i<N. A i)"
         proof (induct N)
-          case 0 thus ?case by simp
+          case 0 thus ?case using pos[unfolded positive_def] by simp
         next
           case (Suc n)
           have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)"
@@ -465,30 +373,25 @@
             by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE)
           thus ?thesis by simp
         qed
-      finally show "(\<Sum>\<^isub>\<infinity> n. measure M (A n)) = measure M (\<Union>i. A i)" .
+      finally show "(\<Sum>n. measure M (A n)) = measure M (\<Union>i. A i)" .
     qed
   qed
 qed
 
 lemma (in measure_space) measure_setsum_split:
-  assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
-  assumes "(\<Union>i \<in> r. b i) = space M"
-  assumes "disjoint_family_on b r"
-  shows "\<mu> a = (\<Sum> i \<in> r. \<mu> (a \<inter> (b i)))"
+  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
+  assumes "(\<Union>i\<in>S. B i) = space M"
+  assumes "disjoint_family_on B S"
+  shows "\<mu> A = (\<Sum>i\<in>S. \<mu> (A \<inter> (B i)))"
 proof -
-  have *: "\<mu> a = \<mu> (\<Union>i \<in> r. a \<inter> b i)"
+  have *: "\<mu> A = \<mu> (\<Union>i\<in>S. A \<inter> B i)"
     using assms by auto
   show ?thesis unfolding *
-  proof (rule measure_finitely_additive''[symmetric])
-    show "finite r" using `finite r` by auto
-    { fix i assume "i \<in> r"
-      hence "b i \<in> sets M" using br_in_M by auto
-      thus "a \<inter> b i \<in> sets M" using `a \<in> sets M` by auto
-    }
-    show "disjoint_family_on (\<lambda>i. a \<inter> b i) r"
-      using `disjoint_family_on b r`
+  proof (rule measure_setsum[symmetric])
+    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
+      using `disjoint_family_on B S`
       unfolding disjoint_family_on_def by auto
-  qed
+  qed (insert assms, auto)
 qed
 
 lemma (in measure_space) measure_subadditive:
@@ -506,7 +409,7 @@
 lemma (in measure_space) measure_eq_0:
   assumes "N \<in> sets M" and "\<mu> N = 0" and "K \<subseteq> N" and "K \<in> sets M"
   shows "\<mu> K = 0"
-using measure_mono[OF assms(3,4,1)] assms(2) by auto
+  using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto
 
 lemma (in measure_space) measure_finitely_subadditive:
   assumes "finite I" "A ` I \<subseteq> sets M"
@@ -523,35 +426,38 @@
 
 lemma (in measure_space) measure_countably_subadditive:
   assumes "range f \<subseteq> sets M"
-  shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
+  shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>i. \<mu> (f i))"
 proof -
   have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)"
     unfolding UN_disjointed_eq ..
-  also have "\<dots> = (\<Sum>\<^isub>\<infinity> i. \<mu> (disjointed f i))"
+  also have "\<dots> = (\<Sum>i. \<mu> (disjointed f i))"
     using range_disjointed_sets[OF assms] measure_countably_additive
     by (simp add:  disjoint_family_disjointed comp_def)
-  also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
-  proof (rule psuminf_le, rule measure_mono)
-    fix i show "disjointed f i \<subseteq> f i" by (rule disjointed_subset)
-    show "f i \<in> sets M" "disjointed f i \<in> sets M"
-      using assms range_disjointed_sets[OF assms] by auto
-  qed
+  also have "\<dots> \<le> (\<Sum>i. \<mu> (f i))"
+    using range_disjointed_sets[OF assms] assms
+    by (auto intro!: suminf_le_pos measure_mono positive_measure disjointed_subset)
   finally show ?thesis .
 qed
 
 lemma (in measure_space) measure_UN_eq_0:
-  assumes "\<And> i :: nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M"
+  assumes "\<And>i::nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M"
   shows "\<mu> (\<Union> i. N i) = 0"
-using measure_countably_subadditive[OF assms(2)] assms(1) by auto
+proof -
+  have "0 \<le> \<mu> (\<Union> i. N i)" using assms by auto
+  moreover have "\<mu> (\<Union> i. N i) \<le> 0"
+    using measure_countably_subadditive[OF assms(2)] assms(1) by simp
+  ultimately show ?thesis by simp
+qed
 
 lemma (in measure_space) measure_inter_full_set:
-  assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
+  assumes "S \<in> sets M" "T \<in> sets M" and fin: "\<mu> (T - S) \<noteq> \<infinity>"
   assumes T: "\<mu> T = \<mu> (space M)"
   shows "\<mu> (S \<inter> T) = \<mu> S"
 proof (rule antisym)
   show " \<mu> (S \<inter> T) \<le> \<mu> S"
     using assms by (auto intro!: measure_mono)
 
+  have pos: "0 \<le> \<mu> (T - S)" using assms by auto
   show "\<mu> S \<le> \<mu> (S \<inter> T)"
   proof (rule ccontr)
     assume contr: "\<not> ?thesis"
@@ -560,7 +466,7 @@
     also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
       using assms by (auto intro!: measure_subadditive)
     also have "\<dots> < \<mu> (T - S) + \<mu> S"
-      by (rule pextreal_less_add[OF not_\<omega>]) (insert contr, auto)
+      using fin contr pos by (intro extreal_less_add) auto
     also have "\<dots> = \<mu> (T \<union> S)"
       using assms by (subst measure_additive) auto
     also have "\<dots> \<le> \<mu> (space M)"
@@ -572,11 +478,11 @@
 lemma measure_unique_Int_stable:
   fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
   assumes "Int_stable E"
-  and A: "range A \<subseteq> sets E" "A \<up> space E"
+  and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E"
   and M: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<mu>\<rparr>" (is "measure_space ?M")
   and N: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<nu>\<rparr>" (is "measure_space ?N")
   and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X"
-  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+  and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
   assumes "X \<in> sets (sigma E)"
   shows "\<mu> X = \<nu> X"
 proof -
@@ -585,9 +491,9 @@
     where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \<mu>" by (simp_all add: M)
   interpret N: measure_space ?N
     where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \<nu>" by (simp_all add: N)
-  { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<omega>"
+  { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<infinity>"
     then have [intro]: "F \<in> sets (sigma E)" by auto
-    have "\<nu> F \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` `F \<in> sets E` eq by simp
+    have "\<nu> F \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` `F \<in> sets E` eq by simp
     interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>"
     proof (rule dynkin_systemI, simp_all)
       fix A assume "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
@@ -602,14 +508,14 @@
         and [intro]: "F \<inter> A \<in> sets (sigma E)"
         using `F \<in> sets E` M.sets_into_space by auto
       have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: N.measure_mono)
-      then have "\<nu> (F \<inter> A) \<noteq> \<omega>" using `\<nu> F \<noteq> \<omega>` by auto
+      then have "\<nu> (F \<inter> A) \<noteq> \<infinity>" using `\<nu> F \<noteq> \<infinity>` by auto
       have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono)
-      then have "\<mu> (F \<inter> A) \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` by auto
+      then have "\<mu> (F \<inter> A) \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` by auto
       then have "\<mu> (F \<inter> (space (sigma E) - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding **
         using `F \<inter> A \<in> sets (sigma E)` by (auto intro!: M.measure_Diff)
       also have "\<dots> = \<nu> F - \<nu> (F \<inter> A)" using eq `F \<in> sets E` * by simp
       also have "\<dots> = \<nu> (F \<inter> (space (sigma E) - A))" unfolding **
-        using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<omega>` by (auto intro!: N.measure_Diff[symmetric])
+        using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<infinity>` by (auto intro!: N.measure_Diff[symmetric])
       finally show "space E - A \<in> sets (sigma E) \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))"
         using * by auto
     next
@@ -630,15 +536,13 @@
     have "\<And>D. D \<in> sets (sigma E) \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)"
       by (subst (asm) *) auto }
   note * = this
-  { fix i have "\<mu> (A i \<inter> X) = \<nu> (A i \<inter> X)"
+  let "?A i" = "A i \<inter> X"
+  have A': "range ?A \<subseteq> sets (sigma E)" "incseq ?A"
+    using A(1,2) `X \<in> sets (sigma E)` by (auto simp: incseq_def)
+  { fix i have "\<mu> (?A i) = \<nu> (?A i)"
       using *[of "A i" X] `X \<in> sets (sigma E)` A finite by auto }
-  moreover
-  have "(\<lambda>i. A i \<inter> X) \<up> X"
-    using `X \<in> sets (sigma E)` M.sets_into_space A
-    by (auto simp: isoton_def)
-  then have "(\<lambda>i. \<mu> (A i \<inter> X)) \<up> \<mu> X" "(\<lambda>i. \<nu> (A i \<inter> X)) \<up> \<nu> X"
-    using `X \<in> sets (sigma E)` A by (auto intro!: M.measure_up N.measure_up M.Int simp: subset_eq)
-  ultimately show ?thesis by (simp add: isoton_def)
+  with M.continuity_from_below[OF A'] N.continuity_from_below[OF A']
+  show ?thesis using A(3) `X \<in> sets (sigma E)` by auto
 qed
 
 section "@{text \<mu>}-null sets"
@@ -650,10 +554,10 @@
   shows "N \<union> N' \<in> null_sets"
 proof (intro conjI CollectI)
   show "N \<union> N' \<in> sets M" using assms by auto
-  have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'"
+  then have "0 \<le> \<mu> (N \<union> N')" by simp
+  moreover have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'"
     using assms by (intro measure_subadditive) auto
-  then show "\<mu> (N \<union> N') = 0"
-    using assms by auto
+  ultimately show "\<mu> (N \<union> N') = 0" using assms by auto
 qed
 
 lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))"
@@ -669,11 +573,11 @@
   shows "(\<Union>i. N i) \<in> null_sets"
 proof (intro conjI CollectI)
   show "(\<Union>i. N i) \<in> sets M" using assms by auto
-  have "\<mu> (\<Union>i. N i) \<le> (\<Sum>\<^isub>\<infinity> n. \<mu> (N (Countable.from_nat n)))"
+  then have "0 \<le> \<mu> (\<Union>i. N i)" by simp
+  moreover have "\<mu> (\<Union>i. N i) \<le> (\<Sum>n. \<mu> (N (Countable.from_nat n)))"
     unfolding UN_from_nat[of N]
     using assms by (intro measure_countably_subadditive) auto
-  then show "\<mu> (\<Union>i. N i) = 0"
-    using assms by auto
+  ultimately show "\<mu> (\<Union>i. N i) = 0" using assms by auto
 qed
 
 lemma (in measure_space) null_sets_finite_UN:
@@ -681,10 +585,10 @@
   shows "(\<Union>i\<in>S. A i) \<in> null_sets"
 proof (intro CollectI conjI)
   show "(\<Union>i\<in>S. A i) \<in> sets M" using assms by (intro finite_UN) auto
-  have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))"
+  then have "0 \<le> \<mu> (\<Union>i\<in>S. A i)" by simp
+  moreover have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))"
     using assms by (intro measure_finitely_subadditive) auto
-  then show "\<mu> (\<Union>i\<in>S. A i) = 0"
-    using assms by auto
+  ultimately show "\<mu> (\<Union>i\<in>S. A i) = 0" using assms by auto
 qed
 
 lemma (in measure_space) null_set_Int1:
@@ -731,6 +635,23 @@
   almost_everywhere :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "AE " 10) where
   "almost_everywhere P \<longleftrightarrow> (\<exists>N\<in>null_sets. { x \<in> space M. \<not> P x } \<subseteq> N)"
 
+syntax
+  "_almost_everywhere" :: "'a \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
+
+translations
+  "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)"
+
+lemma (in measure_space) AE_cong_measure:
+  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
+  shows "(AE x in N. P x) \<longleftrightarrow> (AE x. P x)"
+proof -
+  interpret N: measure_space N
+    by (rule measure_space_cong) fact+
+  show ?thesis
+    unfolding N.almost_everywhere_def almost_everywhere_def
+    by (auto simp: assms)
+qed
+
 lemma (in measure_space) AE_I':
   "N \<in> null_sets \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x. P x)"
   unfolding almost_everywhere_def by auto
@@ -741,13 +662,19 @@
 proof
   assume "AE x. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "\<mu> N = 0"
     unfolding almost_everywhere_def by auto
+  have "0 \<le> \<mu> ?P" using assms by simp
   moreover have "\<mu> ?P \<le> \<mu> N"
     using assms N(1,2) by (auto intro: measure_mono)
-  ultimately show "?P \<in> null_sets" using assms by auto
+  ultimately have "\<mu> ?P = 0" unfolding `\<mu> N = 0` by auto
+  then show "?P \<in> null_sets" using assms by simp
 next
   assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I')
 qed
 
+lemma (in measure_space) AE_iff_measurable:
+  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x. P x) \<longleftrightarrow> \<mu> N = 0"
+  using AE_iff_null_set[of P] by simp
+
 lemma (in measure_space) AE_True[intro, simp]: "AE x. True"
   unfolding almost_everywhere_def by auto
 
@@ -760,13 +687,9 @@
   assumes "AE x. P x" "{x\<in>space M. P x} \<in> sets M"
   shows "\<mu> {x\<in>space M. \<not> P x} = 0" (is "\<mu> ?P = 0")
 proof -
-  obtain A where A: "?P \<subseteq> A" "A \<in> sets M" "\<mu> A = 0"
-    using assms by (auto elim!: AE_E)
-  have "?P = space M - {x\<in>space M. P x}" by auto
-  then have "?P \<in> sets M" using assms by auto
-  with assms `?P \<subseteq> A` `A \<in> sets M` have "\<mu> ?P \<le> \<mu> A"
-    by (auto intro!: measure_mono)
-  then show "\<mu> ?P = 0" using A by simp
+  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}"
+    by auto
+  with AE_iff_null_set[of P] assms show ?thesis by auto
 qed
 
 lemma (in measure_space) AE_I:
@@ -788,8 +711,10 @@
 
   show ?thesis
   proof (intro AE_I)
-    show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B
-      using measure_subadditive[of A B] by auto
+    have "0 \<le> \<mu> (A \<union> B)" using A B by auto
+    moreover have "\<mu> (A \<union> B) \<le> 0"
+      using measure_subadditive[of A B] A B by auto
+    ultimately show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B by auto
     show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
       using P imp by auto
   qed
@@ -818,8 +743,8 @@
   "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x. P x) \<longleftrightarrow> (AE x. Q x)"
   by auto
 
-lemma (in measure_space) all_AE_countable:
-  "(\<forall>i::'i::countable. AE x. P i x) \<longleftrightarrow> (AE x. \<forall>i. P i x)"
+lemma (in measure_space) AE_all_countable:
+  "(AE x. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x. P i x)"
 proof
   assume "\<forall>i. AE x. P i x"
   from this[unfolded almost_everywhere_def Bex_def, THEN choice]
@@ -833,6 +758,10 @@
     unfolding almost_everywhere_def by auto
 qed auto
 
+lemma (in measure_space) AE_finite_all:
+  assumes f: "finite S" shows "(AE x. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x. P i x)"
+  using f by induct auto
+
 lemma (in measure_space) restricted_measure_space:
   assumes "S \<in> sets M"
   shows "measure_space (restricted_space S)"
@@ -840,7 +769,7 @@
   unfolding measure_space_def measure_space_axioms_def
 proof safe
   show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
-  show "measure ?r {} = 0" by simp
+  show "positive ?r (measure ?r)" using `S \<in> sets M` by (auto simp: positive_def)
 
   show "countably_additive ?r (measure ?r)"
     unfolding countably_additive_def
@@ -848,13 +777,38 @@
     fix A :: "nat \<Rightarrow> 'a set"
     assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A"
     from restriction_in_sets[OF assms *[simplified]] **
-    show "(\<Sum>\<^isub>\<infinity> n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
+    show "(\<Sum>n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
       using measure_countably_additive by simp
   qed
 qed
 
+lemma (in measure_space) AE_restricted:
+  assumes "A \<in> sets M"
+  shows "(AE x in restricted_space A. P x) \<longleftrightarrow> (AE x. x \<in> A \<longrightarrow> P x)"
+proof -
+  interpret R: measure_space "restricted_space A"
+    by (rule restricted_measure_space[OF `A \<in> sets M`])
+  show ?thesis
+  proof
+    assume "AE x in restricted_space A. P x"
+    from this[THEN R.AE_E] guess N' .
+    then obtain N where "{x \<in> A. \<not> P x} \<subseteq> A \<inter> N" "\<mu> (A \<inter> N) = 0" "N \<in> sets M"
+      by auto
+    moreover then have "{x \<in> space M. \<not> (x \<in> A \<longrightarrow> P x)} \<subseteq> A \<inter> N"
+      using `A \<in> sets M` sets_into_space by auto
+    ultimately show "AE x. x \<in> A \<longrightarrow> P x"
+      using `A \<in> sets M` by (auto intro!: AE_I[where N="A \<inter> N"])
+  next
+    assume "AE x. x \<in> A \<longrightarrow> P x"
+    from this[THEN AE_E] guess N .
+    then show "AE x in restricted_space A. P x"
+      using null_set_Int1[OF _ `A \<in> sets M`] `A \<in> sets M`[THEN sets_into_space]
+      by (auto intro!: R.AE_I[where N="A \<inter> N"] simp: subset_eq)
+  qed
+qed
+
 lemma (in measure_space) measure_space_subalgebra:
-  assumes "sigma_algebra N" and [simp]: "sets N \<subseteq> sets M" "space N = space M"
+  assumes "sigma_algebra N" and "sets N \<subseteq> sets M" "space N = space M"
   and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X"
   shows "measure_space N"
 proof -
@@ -864,13 +818,26 @@
     from `sets N \<subseteq> sets M` have "\<And>A. range A \<subseteq> sets N \<Longrightarrow> range A \<subseteq> sets M" by blast
     then show "countably_additive N (measure N)"
       by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq)
-  qed simp
+    show "positive N (measure_space.measure N)"
+      using assms(2) by (auto simp add: positive_def)
+  qed
+qed
+
+lemma (in measure_space) AE_subalgebra:
+  assumes ae: "AE x in N. P x"
+  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
+  and sa: "sigma_algebra N"
+  shows "AE x. P x"
+proof -
+  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
+  from ae[THEN N.AE_E] guess N .
+  with N show ?thesis unfolding almost_everywhere_def by auto
 qed
 
 section "@{text \<sigma>}-finite Measures"
 
 locale sigma_finite_measure = measure_space +
-  assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+  assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)"
 
 lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
   assumes "S \<in> sets M"
@@ -881,9 +848,9 @@
   show "measure_space ?r" using restricted_measure_space[OF assms] .
 next
   obtain A :: "nat \<Rightarrow> 'a set" where
-      "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+      "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
     using sigma_finite by auto
-  show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<omega>)"
+  show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<infinity>)"
   proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
     fix i
     show "A i \<inter> S \<in> sets ?r"
@@ -897,8 +864,7 @@
     fix i
     have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)"
       using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono)
-    also have "\<dots> < \<omega>" using `\<mu> (A i) \<noteq> \<omega>` by (auto simp: pextreal_less_\<omega>)
-    finally show "measure ?r (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pextreal_less_\<omega>)
+    then show "measure ?r (A i \<inter> S) \<noteq> \<infinity>" using `\<mu> (A i) \<noteq> \<infinity>` by auto
   qed
 qed
 
@@ -909,7 +875,7 @@
   interpret M': measure_space M' by (intro measure_space_cong cong)
   from sigma_finite guess A .. note A = this
   then have "\<And>i. A i \<in> sets M" by auto
-  with A have fin: "(\<forall>i. measure M' (A i) \<noteq> \<omega>)" using cong by auto
+  with A have fin: "\<forall>i. measure M' (A i) \<noteq> \<infinity>" using cong by auto
   show ?thesis
     apply default
     using A fin cong by auto
@@ -917,30 +883,30 @@
 
 lemma (in sigma_finite_measure) disjoint_sigma_finite:
   "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
-    (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
+    (\<forall>i. \<mu> (A i) \<noteq> \<infinity>) \<and> disjoint_family A"
 proof -
   obtain A :: "nat \<Rightarrow> 'a set" where
     range: "range A \<subseteq> sets M" and
     space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+    measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
     using sigma_finite by auto
   note range' = range_disjointed_sets[OF range] range
   { fix i
     have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
       using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
-    then have "\<mu> (disjointed A i) \<noteq> \<omega>"
+    then have "\<mu> (disjointed A i) \<noteq> \<infinity>"
       using measure[of i] by auto }
   with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
   show ?thesis by (auto intro!: exI[of _ "disjointed A"])
 qed
 
 lemma (in sigma_finite_measure) sigma_finite_up:
-  "\<exists>F. range F \<subseteq> sets M \<and> F \<up> space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<omega>)"
+  "\<exists>F. range F \<subseteq> sets M \<and> incseq F \<and> (\<Union>i. F i) = space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<infinity>)"
 proof -
   obtain F :: "nat \<Rightarrow> 'a set" where
-    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<omega>"
+    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
     using sigma_finite by auto
-  then show ?thesis unfolding isoton_def
+  then show ?thesis
   proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI)
     from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
     then show "(\<Union>n. \<Union> i\<le>n. F i) = space M"
@@ -949,16 +915,16 @@
     fix n
     have "\<mu> (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. \<mu> (F i))" using F
       by (auto intro!: measure_finitely_subadditive)
-    also have "\<dots> < \<omega>"
-      using F by (auto simp: setsum_\<omega>)
-    finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<omega>" by simp
-  qed force+
+    also have "\<dots> < \<infinity>"
+      using F by (auto simp: setsum_Pinfty)
+    finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
+  qed (force simp: incseq_def)+
 qed
 
 section {* Measure preserving *}
 
 definition "measure_preserving A B =
-    {f \<in> measurable A B. (\<forall>y \<in> sets B. measure A (f -` y \<inter> space A) = measure B y)}"
+    {f \<in> measurable A B. (\<forall>y \<in> sets B. measure B y = measure A (f -` y \<inter> space A))}"
 
 lemma measure_preservingI[intro?]:
   assumes "f \<in> measurable A B"
@@ -974,7 +940,8 @@
   interpret M': sigma_algebra M' by fact
   show ?thesis
   proof
-    show "measure M' {} = 0" using T by (force simp: measure_preserving_def)
+    show "positive M' (measure M')" using T
+      by (auto simp: measure_preserving_def positive_def measurable_sets)
 
     show "countably_additive M' (measure M')"
     proof (intro countably_additiveI)
@@ -986,7 +953,7 @@
         using * by blast
       moreover have **: "disjoint_family (\<lambda>i. T -` A i \<inter> space M)"
         using `disjoint_family A` by (auto simp: disjoint_family_on_def)
-      ultimately show "(\<Sum>\<^isub>\<infinity> i. measure M' (A i)) = measure M' (\<Union>i. A i)"
+      ultimately show "(\<Sum>i. measure M' (A i)) = measure M' (\<Union>i. A i)"
         using measure_countably_additive[OF _ **] A T
         by (auto simp: comp_def vimage_UN measure_preserving_def)
     qed
@@ -1009,13 +976,13 @@
 lemma measure_unique_Int_stable_vimage:
   fixes A :: "nat \<Rightarrow> 'a set"
   assumes E: "Int_stable E"
-  and A: "range A \<subseteq> sets E" "A \<up> space E" "\<And>i. measure M (A i) \<noteq> \<omega>"
+  and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure M (A i) \<noteq> \<infinity>"
   and N: "measure_space N" "T \<in> measurable N M"
   and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M"
   and eq: "\<And>X. X \<in> sets E \<Longrightarrow> measure M X = measure N (T -` X \<inter> space N)"
   assumes X: "X \<in> sets (sigma E)"
   shows "measure M X = measure N (T -` X \<inter> space N)"
-proof (rule measure_unique_Int_stable[OF E A(1,2) _ _ eq _ X])
+proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X])
   interpret M: measure_space M by fact
   interpret N: measure_space N by fact
   let "?T X" = "T -` X \<inter> space N"
@@ -1028,12 +995,12 @@
     show "T \<in> measure_preserving N ?E"
       using `T \<in> measurable N M` by (auto simp: M measurable_def measure_preserving_def)
   qed
-  show "\<And>i. M.\<mu> (A i) \<noteq> \<omega>" by fact
+  show "\<And>i. M.\<mu> (A i) \<noteq> \<infinity>" by fact
 qed
 
 lemma (in measure_space) measure_preserving_Int_stable:
   fixes A :: "nat \<Rightarrow> 'a set"
-  assumes E: "Int_stable E" "range A \<subseteq> sets E" "A \<up> space E" "\<And>i. measure E (A i) \<noteq> \<omega>"
+  assumes E: "Int_stable E" "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure E (A i) \<noteq> \<infinity>"
   and N: "measure_space (sigma E)"
   and T: "T \<in> measure_preserving M E"
   shows "T \<in> measure_preserving M (sigma E)"
@@ -1046,7 +1013,7 @@
   show "\<mu> (T -` X \<inter> space M) = E.\<mu> X"
   proof (rule measure_unique_Int_stable_vimage[symmetric])
     show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)"
-      "\<And>i. E.\<mu> (A i) \<noteq> \<omega>" using E by auto
+      "\<And>i. E.\<mu> (A i) \<noteq> \<infinity>" using E by auto
     show "measure_space M" by default
   next
     fix X assume "X \<in> sets E" then show "E.\<mu> X = \<mu> (T -` X \<inter> space M)"
@@ -1057,28 +1024,28 @@
 section "Real measure values"
 
 lemma (in measure_space) real_measure_Union:
-  assumes finite: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+  assumes finite: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
   and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
   shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
   unfolding measure_additive[symmetric, OF measurable]
-  using finite by (auto simp: real_of_pextreal_add)
+  using measurable(1,2)[THEN positive_measure]
+  using finite by (cases rule: extreal2_cases[of "\<mu> A" "\<mu> B"]) auto
 
 lemma (in measure_space) real_measure_finite_Union:
   assumes measurable:
     "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S"
-  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<omega>"
+  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>"
   shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-  using real_of_pextreal_setsum[of S, OF finite]
-        measure_finitely_additive''[symmetric, OF measurable]
-  by simp
+  using finite measurable(2)[THEN positive_measure]
+  by (force intro!: setsum_real_of_extreal[symmetric]
+            simp: measure_setsum[OF measurable, symmetric])
 
 lemma (in measure_space) real_measure_Diff:
-  assumes finite: "\<mu> A \<noteq> \<omega>"
+  assumes finite: "\<mu> A \<noteq> \<infinity>"
   and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
   shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
 proof -
-  have "\<mu> (A - B) \<le> \<mu> A"
-    "\<mu> B \<le> \<mu> A"
+  have "\<mu> (A - B) \<le> \<mu> A" "\<mu> B \<le> \<mu> A"
     using measurable by (auto intro!: measure_mono)
   hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)"
     using measurable finite by (rule_tac real_measure_Union) auto
@@ -1087,117 +1054,155 @@
 
 lemma (in measure_space) real_measure_UNION:
   assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
-  assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+  assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
 proof -
-  have *: "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
-    using measure_countably_additive[OF measurable] by (simp add: comp_def)
-  hence "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) \<noteq> \<omega>" using finite by simp
-  from psuminf_imp_suminf[OF this]
-  show ?thesis using * by simp
+  have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto
+  with summable_sums[OF summable_extreal_pos, of "\<lambda>i. \<mu> (A i)"]
+     measure_countably_additive[OF measurable]
+  have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp
+  moreover
+  { fix i
+    have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)"
+      using measurable by (auto intro!: measure_mono)
+    moreover have "0 \<le> \<mu> (A i)" using measurable by auto
+    ultimately have "\<mu> (A i) = extreal (real (\<mu> (A i)))"
+      using finite by (cases "\<mu> (A i)") auto }
+  moreover
+  have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto
+  then have "\<mu> (\<Union>i. A i) = extreal (real (\<mu> (\<Union>i. A i)))"
+    using finite by (cases "\<mu> (\<Union>i. A i)") auto
+  ultimately show ?thesis
+    unfolding sums_extreal[symmetric] by simp
 qed
 
 lemma (in measure_space) real_measure_subadditive:
   assumes measurable: "A \<in> sets M" "B \<in> sets M"
-  and fin: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+  and fin: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
   shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
 proof -
-  have "real (\<mu> (A \<union> B)) \<le> real (\<mu> A + \<mu> B)"
-    using measure_subadditive[OF measurable] fin by (auto intro!: real_of_pextreal_mono)
-  also have "\<dots> = real (\<mu> A) + real (\<mu> B)"
-    using fin by (simp add: real_of_pextreal_add)
-  finally show ?thesis .
-qed
-
-lemma (in measure_space) real_measure_countably_subadditive:
-  assumes "range f \<subseteq> sets M" and "(\<Sum>\<^isub>\<infinity> i. \<mu> (f i)) \<noteq> \<omega>"
-  shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
-proof -
-  have "real (\<mu> (\<Union>i. f i)) \<le> real (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
-    using assms by (auto intro!: real_of_pextreal_mono measure_countably_subadditive)
-  also have "\<dots> = (\<Sum> i. real (\<mu> (f i)))"
-    using assms by (auto intro!: sums_unique psuminf_imp_suminf)
-  finally show ?thesis .
+  have "0 \<le> \<mu> (A \<union> B)" using measurable by auto
+  then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
+    using measure_subadditive[OF measurable] fin
+    by (cases rule: extreal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
 qed
 
 lemma (in measure_space) real_measure_setsum_singleton:
   assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+  and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
   shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
-  using measure_finite_singleton[OF S] fin by (simp add: real_of_pextreal_setsum)
+  using measure_finite_singleton[OF S] fin
+  using positive_measure[OF S(2)]
+  by (force intro!: setsum_real_of_extreal[symmetric])
 
 lemma (in measure_space) real_continuity_from_below:
-  assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)" and "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
-proof (rule SUP_eq_LIMSEQ[THEN iffD1])
-  { fix n have "\<mu> (A n) \<le> \<mu> (\<Union>i. A i)"
-      using A by (auto intro!: measure_mono)
-    hence "\<mu> (A n) \<noteq> \<omega>" using assms by auto }
-  note this[simp]
+proof -
+  have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
+  then have "extreal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
+    using fin by (auto intro: extreal_real')
+  then show ?thesis
+    using continuity_from_below_Lim[OF A]
+    by (intro lim_real_of_extreal) simp
+qed
 
-  show "mono (\<lambda>i. real (\<mu> (A i)))" unfolding mono_iff_le_Suc using A
-    by (auto intro!: real_of_pextreal_mono measure_mono)
-
-  show "(SUP n. Real (real (\<mu> (A n)))) = Real (real (\<mu> (\<Union>i. A i)))"
-    using continuity_from_below[OF A(1), OF A(2)]
-    using assms by (simp add: Real_real)
-qed simp_all
+lemma (in measure_space) continuity_from_above_Lim:
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)"
+  using LIMSEQ_extreal_INFI[OF measure_decseq, OF A]
+  using continuity_from_above[OF A fin] by simp
 
 lemma (in measure_space) real_continuity_from_above:
-  assumes A: "range A \<subseteq> sets M"
-  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
-  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
   shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
-proof (rule INF_eq_LIMSEQ[THEN iffD1])
-  { fix n have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A n)"
-      using A by (auto intro!: measure_mono)
-    hence "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using assms by auto }
-  note this[simp]
+proof -
+  have "0 \<le> \<mu> (\<Inter>i. A i)" using A by auto
+  moreover
+  have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
+    using A by (auto intro!: measure_mono)
+  ultimately have "extreal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
+    using fin by (auto intro: extreal_real')
+  then show ?thesis
+    using continuity_from_above_Lim[OF A fin]
+    by (intro lim_real_of_extreal) simp
+qed
 
-  show "mono (\<lambda>i. - real (\<mu> (A i)))" unfolding mono_iff_le_Suc using assms
-    by (auto intro!: real_of_pextreal_mono measure_mono)
-
-  show "(INF n. Real (real (\<mu> (A n)))) =
-    Real (real (\<mu> (\<Inter>i. A i)))"
-    using continuity_from_above[OF A, OF mono_Suc finite]
-    using assms by (simp add: Real_real)
-qed simp_all
+lemma (in measure_space) real_measure_countably_subadditive:
+  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. \<mu> (A i)) \<noteq> \<infinity>"
+  shows "real (\<mu> (\<Union>i. A i)) \<le> (\<Sum>i. real (\<mu> (A i)))"
+proof -
+  { fix i
+    have "0 \<le> \<mu> (A i)" using A by auto
+    moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto
+    ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto }
+  moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
+  ultimately have "extreal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. extreal (real (\<mu> (A i))))"
+    using measure_countably_subadditive[OF A] by (auto simp: extreal_real)
+  also have "\<dots> = extreal (\<Sum>i. real (\<mu> (A i)))"
+    using A
+    by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin)
+  finally show ?thesis by simp
+qed
 
 locale finite_measure = measure_space +
-  assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<omega>"
+  assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<infinity>"
 
 sublocale finite_measure < sigma_finite_measure
 proof
-  show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+  show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)"
     using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
 qed
 
 lemma (in finite_measure) finite_measure[simp, intro]:
   assumes "A \<in> sets M"
-  shows "\<mu> A \<noteq> \<omega>"
+  shows "\<mu> A \<noteq> \<infinity>"
 proof -
   from `A \<in> sets M` have "A \<subseteq> space M"
     using sets_into_space by blast
-  hence "\<mu> A \<le> \<mu> (space M)"
+  then have "\<mu> A \<le> \<mu> (space M)"
     using assms top by (rule measure_mono)
-  also have "\<dots> < \<omega>"
-    using finite_measure_of_space unfolding pextreal_less_\<omega> .
-  finally show ?thesis unfolding pextreal_less_\<omega> .
+  then show ?thesis
+    using finite_measure_of_space by auto
 qed
 
+lemma (in finite_measure) measure_not_inf:
+  assumes A: "A \<in> sets M"
+  shows "\<bar>\<mu> A\<bar> \<noteq> \<infinity>"
+  using finite_measure[OF A] positive_measure[OF A] by auto
+
+definition (in finite_measure)
+  "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)"
+
+lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = extreal (\<mu>' A)"
+  using measure_not_inf[of A] by (auto simp: \<mu>'_def)
+
+lemma (in finite_measure) positive_measure': "0 \<le> \<mu>' A"
+  unfolding \<mu>'_def by (auto simp: real_of_extreal_pos)
+
+lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)"
+proof cases
+  assume "A \<in> sets M"
+  moreover then have "\<mu> A \<le> \<mu> (space M)"
+    using sets_into_space by (auto intro!: measure_mono)
+  ultimately show ?thesis
+    using measure_not_inf[of A] measure_not_inf[of "space M"]
+    by (auto simp: \<mu>'_def)
+qed (simp add: \<mu>'_def real_of_extreal_pos)
+
 lemma (in finite_measure) restricted_finite_measure:
   assumes "S \<in> sets M"
   shows "finite_measure (restricted_space S)"
     (is "finite_measure ?r")
   unfolding finite_measure_def finite_measure_axioms_def
-proof (safe del: notI)
+proof (intro conjI)
   show "measure_space ?r" using restricted_measure_space[OF assms] .
 next
-  show "measure ?r (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
+  show "measure ?r (space ?r) \<noteq> \<infinity>" using finite_measure[OF `S \<in> sets M`] by auto
 qed
 
 lemma (in measure_space) restricted_to_finite_measure:
-  assumes "S \<in> sets M" "\<mu> S \<noteq> \<omega>"
+  assumes "S \<in> sets M" "\<mu> S \<noteq> \<infinity>"
   shows "finite_measure (restricted_space S)"
 proof -
   have "measure_space (restricted_space S)"
@@ -1207,202 +1212,128 @@
     using assms by auto
 qed
 
-lemma (in finite_measure) real_finite_measure_Diff:
-  assumes "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
-  shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
-  using finite_measure[OF `A \<in> sets M`] by (rule real_measure_Diff[OF _ assms])
-
-lemma (in finite_measure) real_finite_measure_Union:
-  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
-  shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
-  using sets by (auto intro!: real_measure_Union[OF _ _ assms] finite_measure)
+lemma (in finite_measure) finite_measure_Diff:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+  shows "\<mu>' (A - B) = \<mu>' A - \<mu>' B"
+  using sets[THEN finite_measure_eq]
+  using Diff[OF sets, THEN finite_measure_eq]
+  using measure_Diff[OF _ assms] by simp
 
-lemma (in finite_measure) real_finite_measure_finite_Union:
-  assumes "finite S" and dis: "disjoint_family_on A S"
-  and *: "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
-  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-proof (rule real_measure_finite_Union[OF `finite S` _ dis])
-  fix i assume "i \<in> S" from *[OF this] show "A i \<in> sets M" .
-  from finite_measure[OF this] show "\<mu> (A i) \<noteq> \<omega>" .
-qed
-
-lemma (in finite_measure) real_finite_measure_UNION:
-  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
-  shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
-proof (rule real_measure_UNION[OF assms])
-  have "(\<Union>i. A i) \<in> sets M" using measurable(1) by (rule countable_UN)
-  thus "\<mu> (\<Union>i. A i) \<noteq> \<omega>" by (rule finite_measure)
-qed
-
-lemma (in finite_measure) real_measure_mono:
-  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<subseteq> B \<Longrightarrow> real (\<mu> A) \<le> real (\<mu> B)"
-  by (auto intro!: measure_mono real_of_pextreal_mono finite_measure)
-
-lemma (in finite_measure) real_finite_measure_subadditive:
-  assumes measurable: "A \<in> sets M" "B \<in> sets M"
-  shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
-  using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive)
+lemma (in finite_measure) finite_measure_Union:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
+  shows "\<mu>' (A \<union> B) = \<mu>' A + \<mu>' B"
+  using measure_additive[OF assms]
+  using sets[THEN finite_measure_eq]
+  using Un[OF sets, THEN finite_measure_eq]
+  by simp
 
-lemma (in finite_measure) real_finite_measure_countably_subadditive:
-  assumes "range f \<subseteq> sets M" and "summable (\<lambda>i. real (\<mu> (f i)))"
-  shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
-proof (rule real_measure_countably_subadditive[OF assms(1)])
-  have *: "\<And>i. f i \<in> sets M" using assms by auto
-  have "(\<lambda>i. real (\<mu> (f i))) sums (\<Sum>i. real (\<mu> (f i)))"
-    using assms(2) by (rule summable_sums)
-  from suminf_imp_psuminf[OF this]
-  have "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) = Real (\<Sum>i. real (\<mu> (f i)))"
-    using * by (simp add: Real_real finite_measure)
-  thus "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) \<noteq> \<omega>" by simp
-qed
-
-lemma (in finite_measure) real_finite_measure_finite_singelton:
-  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
-proof (rule real_measure_setsum_singleton[OF `finite S`])
-  fix x assume "x \<in> S" thus "{x} \<in> sets M" by (rule *)
-  with finite_measure show "\<mu> {x} \<noteq> \<omega>" .
-qed
+lemma (in finite_measure) finite_measure_finite_Union:
+  assumes S: "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
+  and dis: "disjoint_family_on A S"
+  shows "\<mu>' (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. \<mu>' (A i))"
+  using measure_setsum[OF assms]
+  using finite_UN[of S A, OF S, THEN finite_measure_eq]
+  using S(2)[THEN finite_measure_eq]
+  by simp
 
-lemma (in finite_measure) real_finite_continuity_from_below:
-  assumes "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
-  shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
-  using real_continuity_from_below[OF assms(1), OF assms(2) finite_measure] assms by auto
-
-lemma (in finite_measure) real_finite_continuity_from_above:
-  assumes A: "range A \<subseteq> sets M"
-  and mono_Suc: "\<And>n.  A (Suc n) \<subseteq> A n"
-  shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
-  using real_continuity_from_above[OF A, OF mono_Suc finite_measure] A
-  by auto
-
-lemma (in finite_measure) real_finite_measure_finite_Union':
-  assumes "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
-  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-  using assms real_finite_measure_finite_Union[of S A] by auto
-
-lemma (in finite_measure) finite_measure_compl:
-  assumes s: "s \<in> sets M"
-  shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
-  using measure_compl[OF s, OF finite_measure, OF s] .
-
-lemma (in finite_measure) finite_measure_inter_full_set:
-  assumes "S \<in> sets M" "T \<in> sets M"
-  assumes T: "\<mu> T = \<mu> (space M)"
-  shows "\<mu> (S \<inter> T) = \<mu> S"
-  using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
+lemma (in finite_measure) finite_measure_UNION:
+  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
+  shows "(\<lambda>i. \<mu>' (A i)) sums (\<mu>' (\<Union>i. A i))"
+  using real_measure_UNION[OF A]
+  using countable_UN[OF A(1), THEN finite_measure_eq]
+  using A(1)[THEN subsetD, THEN finite_measure_eq]
   by auto
 
-lemma (in finite_measure) measure_preserving_lift:
-  fixes f :: "'a \<Rightarrow> 'c" and A :: "('c, 'd) measure_space_scheme"
-  assumes "algebra A" "finite_measure (sigma A)" (is "finite_measure ?sA")
-  assumes fmp: "f \<in> measure_preserving M A"
-  shows "f \<in> measure_preserving M (sigma A)"
+lemma (in finite_measure) finite_measure_mono:
+  assumes B: "B \<in> sets M" and "A \<subseteq> B" shows "\<mu>' A \<le> \<mu>' B"
+proof cases
+  assume "A \<in> sets M"
+  from this[THEN finite_measure_eq] B[THEN finite_measure_eq]
+  show ?thesis using measure_mono[OF `A \<subseteq> B` `A \<in> sets M` `B \<in> sets M`] by simp
+next
+  assume "A \<notin> sets M" then show ?thesis
+    using positive_measure'[of B] unfolding \<mu>'_def by auto
+qed
+
+lemma (in finite_measure) finite_measure_subadditive:
+  assumes m: "A \<in> sets M" "B \<in> sets M"
+  shows "\<mu>' (A \<union> B) \<le> \<mu>' A + \<mu>' B"
+  using measure_subadditive[OF m]
+  using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp
+
+lemma (in finite_measure) finite_measure_countably_subadditive:
+  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. \<mu>' (A i))"
+  shows "\<mu>' (\<Union>i. A i) \<le> (\<Sum>i. \<mu>' (A i))"
 proof -
-  interpret sA: finite_measure ?sA by fact
-  interpret A: algebra A by fact
-  show ?thesis using fmp
-  proof (clarsimp simp add: measure_preserving_def)
-    assume fm: "f \<in> measurable M A"
-       and "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = measure A y"
-    then have meq: "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = sA.\<mu> y"
-      by simp
-    have f12: "f \<in> measurable M ?sA"
-      using measurable_subset[OF A.sets_into_space] fm by auto
-    hence ffn: "f \<in> space M \<rightarrow> space A"
-      by (simp add: measurable_def)
-    have "space M \<subseteq> f -` (space A)"
-      by auto (metis PiE ffn)
-    hence fveq [simp]: "(f -` (space A)) \<inter> space M = space M"
-      by blast
-    {
-      fix y
-      assume y: "y \<in> sets ?sA"
-      have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def)
-      also have "\<dots> \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = sA.\<mu> s}"
-      proof (rule A.sigma_property_disjoint, safe)
-        fix x assume "x \<in> sets A" then show "\<mu> (f -` x \<inter> space M) = sA.\<mu> x" by (simp add: meq)
-      next
-        fix s
-        assume eq: "\<mu> (f -` s \<inter> space M) = sA.\<mu> s" and s: "s \<in> ?A"
-        then have s': "s \<in> sets ?sA" by (simp add: sigma_def)
-        show "\<mu> (f -` (space A - s) \<inter> space M) = measure ?sA (space A - s)"
-          using sA.finite_measure_compl[OF s']
-          using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top]
-          by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq)
-      next
-        fix S
-        assume S0: "S 0 = {}"
-           and SSuc: "\<And>n.  S n \<subseteq> S (Suc n)"
-           and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = sA.\<mu> s} \<inter> ?A"
-        hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
-        have eq1: "\<And>i. \<mu> (f -` S i \<inter> space M) = sA.\<mu> (S i)"
-          using rS1 by blast
-        have *: "(\<lambda>n. sA.\<mu> (S n)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
-          by (simp add: eq1)
-        have "(SUP n. ... n) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
-        proof (rule measure_countable_increasing)
-          show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
-            using f12 rS2 by (auto simp add: measurable_def)
-          show "f -` S 0 \<inter> space M = {}" using S0
-            by blast
-          show "\<And>n. f -` S n \<inter> space M \<subseteq> f -` S (Suc n) \<inter> space M"
-            using SSuc by auto
-        qed
-        also have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)"
-          by (simp add: vimage_UN)
-        finally have "(SUP n. sA.\<mu> (S n)) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)" unfolding * .
-        moreover
-        have "(SUP n. sA.\<mu> (S n)) = sA.\<mu> (\<Union>i. S i)"
-          by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc])
-        ultimately
-        show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = sA.\<mu> (\<Union>i. S i)" by simp
-      next
-        fix S :: "nat \<Rightarrow> 'c set"
-        assume dS: "disjoint_family S"
-           and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = sA.\<mu> s} \<inter> ?A"
-        hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
-        have "\<And>i. \<mu> (f -` S i \<inter> space M) = sA.\<mu> (S i)"
-          using rS1 by blast
-        hence *: "(\<lambda>i. sA.\<mu> (S i)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
-          by simp
-        have "psuminf ... = \<mu> (\<Union>i. f -` S i \<inter> space M)"
-          proof (rule measure_countably_additive)
-            show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
-              using f12 rS2 by (auto simp add: measurable_def)
-            show "disjoint_family (\<lambda>i. f -` S i \<inter> space M)" using dS
-              by (auto simp add: disjoint_family_on_def)
-          qed
-        hence "(\<Sum>\<^isub>\<infinity> i. sA.\<mu> (S i)) = \<mu> (\<Union>i. f -` S i \<inter> space M)" unfolding * .
-        with sA.measure_countably_additive [OF rS2 dS]
-        have "\<mu> (\<Union>i. f -` S i \<inter> space M) = sA.\<mu> (\<Union>i. S i)"
-          by simp
-        moreover have "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
-          by (simp add: vimage_UN)
-        ultimately show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = sA.\<mu> (\<Union>i. S i)"
-          by metis
-      qed
-      finally have "sets ?sA \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = sA.\<mu> s}" .
-      hence "\<mu> (f -` y \<inter> space M) = sA.\<mu> y" using y by force
-    }
-    thus "f \<in> measurable M ?sA \<and> (\<forall>y\<in>sets ?sA. \<mu> (f -` y \<inter> space M) = measure A y)"
-      by simp_all (blast intro: f12)
-  qed
+  note A[THEN subsetD, THEN finite_measure_eq, simp]
+  note countable_UN[OF A, THEN finite_measure_eq, simp]
+  from `summable (\<lambda>i. \<mu>' (A i))`
+  have "(\<lambda>i. extreal (\<mu>' (A i))) sums extreal (\<Sum>i. \<mu>' (A i))"
+    by (simp add: sums_extreal) (rule summable_sums)
+  from sums_unique[OF this, symmetric]
+       measure_countably_subadditive[OF A]
+  show ?thesis by simp
 qed
 
+lemma (in finite_measure) finite_measure_finite_singleton:
+  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "\<mu>' S = (\<Sum>x\<in>S. \<mu>' {x})"
+  using real_measure_setsum_singleton[OF assms]
+  using *[THEN finite_measure_eq]
+  using finite_UN[of S "\<lambda>x. {x}", OF assms, THEN finite_measure_eq]
+  by simp
+
+lemma (in finite_measure) finite_continuity_from_below:
+  assumes A: "range A \<subseteq> sets M" and "incseq A"
+  shows "(\<lambda>i. \<mu>' (A i)) ----> \<mu>' (\<Union>i. A i)"
+  using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms
+  using A[THEN subsetD, THEN finite_measure_eq]
+  using countable_UN[OF A, THEN finite_measure_eq]
+  by auto
+
+lemma (in finite_measure) finite_continuity_from_above:
+  assumes A: "range A \<subseteq> sets M" and "decseq A"
+  shows "(\<lambda>n. \<mu>' (A n)) ----> \<mu>' (\<Inter>i. A i)"
+  using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms
+  using A[THEN subsetD, THEN finite_measure_eq]
+  using countable_INT[OF A, THEN finite_measure_eq]
+  by auto
+
+lemma (in finite_measure) finite_measure_compl:
+  assumes S: "S \<in> sets M"
+  shows "\<mu>' (space M - S) = \<mu>' (space M) - \<mu>' S"
+  using measure_compl[OF S, OF finite_measure, OF S]
+  using S[THEN finite_measure_eq]
+  using compl_sets[OF S, THEN finite_measure_eq]
+  using top[THEN finite_measure_eq]
+  by simp
+
+lemma (in finite_measure) finite_measure_inter_full_set:
+  assumes S: "S \<in> sets M" "T \<in> sets M"
+  assumes T: "\<mu>' T = \<mu>' (space M)"
+  shows "\<mu>' (S \<inter> T) = \<mu>' S"
+  using measure_inter_full_set[OF S finite_measure]
+  using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq]
+  using Int[OF S, THEN finite_measure_eq]
+  using S[THEN finite_measure_eq] top[THEN finite_measure_eq]
+  by simp
+
+lemma (in finite_measure) empty_measure'[simp]: "\<mu>' {} = 0"
+  unfolding \<mu>'_def by simp
+
 section "Finite spaces"
 
 locale finite_measure_space = measure_space + finite_sigma_algebra +
-  assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+  assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
 
 lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
-  using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
+  using measure_setsum[of "space M" "\<lambda>i. {i}"]
   by (simp add: sets_eq_Pow disjoint_family_on_def finite_space)
 
 lemma finite_measure_spaceI:
-  assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<omega>"
+  assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \<noteq> \<infinity>"
     and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
-    and "measure M {} = 0"
+    and "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
   shows "finite_measure_space M"
     unfolding finite_measure_space_def finite_measure_space_axioms_def
 proof (intro allI impI conjI)
@@ -1414,7 +1345,7 @@
       using sigma_algebra_Pow[of "space M" "algebra.more M"]
       unfolding * .
     show "finite (space M)" by fact
-    show "positive M (measure M)" unfolding positive_def by fact
+    show "positive M (measure M)" unfolding positive_def using assms by auto
     show "additive M (measure M)" unfolding additive_def using assms by simp
   qed
   then interpret measure_space M .
@@ -1425,19 +1356,20 @@
   qed
   { fix x assume *: "x \<in> space M"
     with add[of "{x}" "space M - {x}"] space
-    show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) }
+    show "\<mu> {x} \<noteq> \<infinity>" by (auto simp: insert_absorb[OF *] Diff_subset) }
 qed
 
 sublocale finite_measure_space \<subseteq> finite_measure
 proof
-  show "\<mu> (space M) \<noteq> \<omega>"
-    unfolding sum_over_space[symmetric] setsum_\<omega>
+  show "\<mu> (space M) \<noteq> \<infinity>"
+    unfolding sum_over_space[symmetric] setsum_Pinfty
     using finite_space finite_single_measure by auto
 qed
 
 lemma finite_measure_space_iff:
   "finite_measure_space M \<longleftrightarrow>
-    finite (space M) \<and> sets M = Pow(space M) \<and> measure M (space M) \<noteq> \<omega> \<and> measure M {} = 0 \<and>
+    finite (space M) \<and> sets M = Pow(space M) \<and> measure M (space M) \<noteq> \<infinity> \<and>
+    measure M {} = 0 \<and> (\<forall>A\<subseteq>space M. 0 \<le> measure M A) \<and>
     (\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> measure M (A \<union> B) = measure M A + measure M B)"
     (is "_ = ?rhs")
 proof (intro iffI)
@@ -1451,29 +1383,31 @@
     by (auto intro!: finite_measure_spaceI)
 qed
 
-lemma psuminf_cmult_indicator:
-  assumes "disjoint_family A" "x \<in> A i"
-  shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
+lemma suminf_cmult_indicator:
+  fixes f :: "nat \<Rightarrow> extreal"
+  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
+  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
 proof -
-  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: pextreal)"
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)"
     using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
-  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: pextreal)"
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: extreal)"
     by (auto simp: setsum_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: pextreal)"
-  proof (rule pextreal_SUPI)
-    fix y :: pextreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+  moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)"
+  proof (rule extreal_SUPI)
+    fix y :: extreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
     from this[of "Suc i"] show "f i \<le> y" by auto
-  qed simp
-  ultimately show ?thesis using `x \<in> A i` unfolding psuminf_def by auto
+  qed (insert assms, simp)
+  ultimately show ?thesis using assms
+    by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def)
 qed
 
-lemma psuminf_indicator:
+lemma suminf_indicator:
   assumes "disjoint_family A"
-  shows "(\<Sum>\<^isub>\<infinity> n. indicator (A n) x) = indicator (\<Union>i. A i) x"
+  shows "(\<Sum>n. indicator (A n) x :: extreal) = indicator (\<Union>i. A i) x"
 proof cases
   assume *: "x \<in> (\<Union>i. A i)"
   then obtain i where "x \<in> A i" by auto
-  from psuminf_cmult_indicator[OF assms, OF this, of "\<lambda>i. 1"]
+  from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
   show ?thesis using * by simp
 qed simp
 
--- a/src/HOL/Probability/Positive_Extended_Real.thy	Mon Mar 14 15:17:10 2011 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2908 +0,0 @@
-(* Author: Johannes Hoelzl, TU Muenchen *)
-
-header {* A type for positive real numbers with infinity *}
-
-theory Positive_Extended_Real
-  imports Complex_Main "~~/src/HOL/Library/Nat_Bijection" Multivariate_Analysis
-begin
-
-lemma (in complete_lattice) Sup_start:
-  assumes *: "\<And>x. f x \<le> f 0"
-  shows "(SUP n. f n) = f 0"
-proof (rule antisym)
-  show "f 0 \<le> (SUP n. f n)" by (rule le_SUPI) auto
-  show "(SUP n. f n) \<le> f 0" by (rule SUP_leI[OF *])
-qed
-
-lemma (in complete_lattice) Inf_start:
-  assumes *: "\<And>x. f 0 \<le> f x"
-  shows "(INF n. f n) = f 0"
-proof (rule antisym)
-  show "(INF n. f n) \<le> f 0" by (rule INF_leI) simp
-  show "f 0 \<le> (INF n. f n)" by (rule le_INFI[OF *])
-qed
-
-lemma (in complete_lattice) Sup_mono_offset:
-  fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
-  assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" and "0 \<le> k"
-  shows "(SUP n . f (k + n)) = (SUP n. f n)"
-proof (rule antisym)
-  show "(SUP n. f (k + n)) \<le> (SUP n. f n)"
-    by (auto intro!: Sup_mono simp: SUPR_def)
-  { fix n :: 'b
-    have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
-    with * have "f n \<le> f (k + n)" by simp }
-  thus "(SUP n. f n) \<le> (SUP n. f (k + n))"
-    by (auto intro!: Sup_mono exI simp: SUPR_def)
-qed
-
-lemma (in complete_lattice) Sup_mono_offset_Suc:
-  assumes *: "\<And>x. f x \<le> f (Suc x)"
-  shows "(SUP n . f (Suc n)) = (SUP n. f n)"
-  unfolding Suc_eq_plus1
-  apply (subst add_commute)
-  apply (rule Sup_mono_offset)
-  by (auto intro!: order.lift_Suc_mono_le[of "op \<le>" "op <" f, OF _ *]) default
-
-lemma (in complete_lattice) Inf_mono_offset:
-  fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
-  assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x" and "0 \<le> k"
-  shows "(INF n . f (k + n)) = (INF n. f n)"
-proof (rule antisym)
-  show "(INF n. f n) \<le> (INF n. f (k + n))"
-    by (auto intro!: Inf_mono simp: INFI_def)
-  { fix n :: 'b
-    have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
-    with * have "f (k + n) \<le> f n" by simp }
-  thus "(INF n. f (k + n)) \<le> (INF n. f n)"
-    by (auto intro!: Inf_mono exI simp: INFI_def)
-qed
-
-lemma (in complete_lattice) isotone_converge:
-  fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y "
-  shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
-proof -
-  have "\<And>n. (SUP m. f (n + m)) = (SUP n. f n)"
-    apply (rule Sup_mono_offset)
-    apply (rule assms)
-    by simp_all
-  moreover
-  { fix n have "(INF m. f (n + m)) = f n"
-      using Inf_start[of "\<lambda>m. f (n + m)"] assms by simp }
-  ultimately show ?thesis by simp
-qed
-
-lemma (in complete_lattice) antitone_converges:
-  fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x"
-  shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
-proof -
-  have "\<And>n. (INF m. f (n + m)) = (INF n. f n)"
-    apply (rule Inf_mono_offset)
-    apply (rule assms)
-    by simp_all
-  moreover
-  { fix n have "(SUP m. f (n + m)) = f n"
-      using Sup_start[of "\<lambda>m. f (n + m)"] assms by simp }
-  ultimately show ?thesis by simp
-qed
-
-lemma (in complete_lattice) lim_INF_le_lim_SUP:
-  fixes f :: "nat \<Rightarrow> 'a"
-  shows "(SUP n. INF m. f (n + m)) \<le> (INF n. SUP m. f (n + m))"
-proof (rule SUP_leI, rule le_INFI)
-  fix i j show "(INF m. f (i + m)) \<le> (SUP m. f (j + m))"
-  proof (cases rule: le_cases)
-    assume "i \<le> j"
-    have "(INF m. f (i + m)) \<le> f (i + (j - i))" by (rule INF_leI) simp
-    also have "\<dots> = f (j + 0)" using `i \<le> j` by auto
-    also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
-    finally show ?thesis .
-  next
-    assume "j \<le> i"
-    have "(INF m. f (i + m)) \<le> f (i + 0)" by (rule INF_leI) simp
-    also have "\<dots> = f (j + (i - j))" using `j \<le> i` by auto
-    also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
-    finally show ?thesis .
-  qed
-qed
-
-text {*
-
-We introduce the the positive real numbers as needed for measure theory.
-
-*}
-
-typedef pextreal = "(Some ` {0::real..}) \<union> {None}"
-  by (rule exI[of _ None]) simp
-
-subsection "Introduce @{typ pextreal} similar to a datatype"
-
-definition "Real x = Abs_pextreal (Some (sup 0 x))"
-definition "\<omega> = Abs_pextreal None"
-
-definition "pextreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
-
-definition "of_pextreal = pextreal_case (\<lambda>x. x) 0"
-
-defs (overloaded)
-  real_of_pextreal_def [code_unfold]: "real == of_pextreal"
-
-lemma pextreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pextreal"
-  unfolding pextreal_def by simp
-
-lemma pextreal_Some_sup[simp]: "Some (sup 0 x) \<in> pextreal"
-  by (simp add: sup_ge1)
-
-lemma pextreal_None[simp]: "None \<in> pextreal"
-  unfolding pextreal_def by simp
-
-lemma Real_inj[simp]:
-  assumes  "0 \<le> x" and "0 \<le> y"
-  shows "Real x = Real y \<longleftrightarrow> x = y"
-  unfolding Real_def assms[THEN sup_absorb2]
-  using assms by (simp add: Abs_pextreal_inject)
-
-lemma Real_neq_\<omega>[simp]:
-  "Real x = \<omega> \<longleftrightarrow> False"
-  "\<omega> = Real x \<longleftrightarrow> False"
-  by (simp_all add: Abs_pextreal_inject \<omega>_def Real_def)
-
-lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
-  unfolding Real_def by (auto simp add: Abs_pextreal_inject intro!: sup_absorb1)
-
-lemma pextreal_cases[case_names preal infinite, cases type: pextreal]:
-  assumes preal: "\<And>r. x = Real r \<Longrightarrow> 0 \<le> r \<Longrightarrow> P" and inf: "x = \<omega> \<Longrightarrow> P"
-  shows P
-proof (cases x rule: pextreal.Abs_pextreal_cases)
-  case (Abs_pextreal y)
-  hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
-    unfolding pextreal_def by auto
-  thus P
-  proof (rule disjE)
-    assume "\<exists>x\<ge>0. y = Some x" then guess x ..
-    thus P by (simp add: preal[of x] Real_def Abs_pextreal(1) sup_absorb2)
-  qed (simp add: \<omega>_def Abs_pextreal(1) inf)
-qed
-
-lemma pextreal_case_\<omega>[simp]: "pextreal_case f i \<omega> = i"
-  unfolding pextreal_case_def by simp
-
-lemma pextreal_case_Real[simp]: "pextreal_case f i (Real x) = (if 0 \<le> x then f x else f 0)"
-proof (cases "0 \<le> x")
-  case True thus ?thesis unfolding pextreal_case_def by (auto intro: theI2)
-next
-  case False
-  moreover have "(THE r. 0 \<le> r \<and> Real 0 = Real r) = 0"
-    by (auto intro!: the_equality)
-  ultimately show ?thesis unfolding pextreal_case_def by (simp add: Real_neg)
-qed
-
-lemma pextreal_case_cancel[simp]: "pextreal_case (\<lambda>c. i) i x = i"
-  by (cases x) simp_all
-
-lemma pextreal_case_split:
-  "P (pextreal_case f i x) = ((x = \<omega> \<longrightarrow> P i) \<and> (\<forall>r\<ge>0. x = Real r \<longrightarrow> P (f r)))"
-  by (cases x) simp_all
-
-lemma pextreal_case_split_asm:
-  "P (pextreal_case f i x) = (\<not> (x = \<omega> \<and> \<not> P i \<or> (\<exists>r. r \<ge> 0 \<and> x = Real r \<and> \<not> P (f r))))"
-  by (cases x) auto
-
-lemma pextreal_case_cong[cong]:
-  assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
-  shows "pextreal_case f i x = pextreal_case f' i' x'"
-  unfolding eq using cong by (cases x') simp_all
-
-lemma real_Real[simp]: "real (Real x) = (if 0 \<le> x then x else 0)"
-  unfolding real_of_pextreal_def of_pextreal_def by simp
-
-lemma Real_real_image:
-  assumes "\<omega> \<notin> A" shows "Real ` real ` A = A"
-proof safe
-  fix x assume "x \<in> A"
-  hence *: "x = Real (real x)"
-    using `\<omega> \<notin> A` by (cases x) auto
-  show "x \<in> Real ` real ` A"
-    using `x \<in> A` by (subst *) (auto intro!: imageI)
-next
-  fix x assume "x \<in> A"
-  thus "Real (real x) \<in> A"
-    using `\<omega> \<notin> A` by (cases x) auto
-qed
-
-lemma real_pextreal_nonneg[simp, intro]: "0 \<le> real (x :: pextreal)"
-  unfolding real_of_pextreal_def of_pextreal_def
-  by (cases x) auto
-
-lemma real_\<omega>[simp]: "real \<omega> = 0"
-  unfolding real_of_pextreal_def of_pextreal_def by simp
-
-lemma pextreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
-
-subsection "@{typ pextreal} is a monoid for addition"
-
-instantiation pextreal :: comm_monoid_add
-begin
-
-definition "0 = Real 0"
-definition "x + y = pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
-
-lemma pextreal_plus[simp]:
-  "Real r + Real p = (if 0 \<le> r then if 0 \<le> p then Real (r + p) else Real r else Real p)"
-  "x + 0 = x"
-  "0 + x = x"
-  "x + \<omega> = \<omega>"
-  "\<omega> + x = \<omega>"
-  by (simp_all add: plus_pextreal_def Real_neg zero_pextreal_def split: pextreal_case_split)
-
-lemma \<omega>_neq_0[simp]:
-  "\<omega> = 0 \<longleftrightarrow> False"
-  "0 = \<omega> \<longleftrightarrow> False"
-  by (simp_all add: zero_pextreal_def)
-
-lemma Real_eq_0[simp]:
-  "Real r = 0 \<longleftrightarrow> r \<le> 0"
-  "0 = Real r \<longleftrightarrow> r \<le> 0"
-  by (auto simp add: Abs_pextreal_inject zero_pextreal_def Real_def sup_real_def)
-
-lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pextreal_def)
-
-instance
-proof
-  fix a :: pextreal
-  show "0 + a = a" by (cases a) simp_all
-
-  fix b show "a + b = b + a"
-    by (cases a, cases b) simp_all
-
-  fix c show "a + b + c = a + (b + c)"
-    by (cases a, cases b, cases c) simp_all
-qed
-end
-
-lemma Real_minus_abs[simp]: "Real (- \<bar>x\<bar>) = 0"
-  by simp
-
-lemma pextreal_plus_eq_\<omega>[simp]: "(a :: pextreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
-  by (cases a, cases b) auto
-
-lemma pextreal_add_cancel_left:
-  "a + b = a + c \<longleftrightarrow> (a = \<omega> \<or> b = c)"
-  by (cases a, cases b, cases c, simp_all, cases c, simp_all)
-
-lemma pextreal_add_cancel_right:
-  "b + a = c + a \<longleftrightarrow> (a = \<omega> \<or> b = c)"
-  by (cases a, cases b, cases c, simp_all, cases c, simp_all)
-
-lemma Real_eq_Real:
-  "Real a = Real b \<longleftrightarrow> (a = b \<or> (a \<le> 0 \<and> b \<le> 0))"
-proof (cases "a \<le> 0 \<or> b \<le> 0")
-  case False with Real_inj[of a b] show ?thesis by auto
-next
-  case True
-  thus ?thesis
-  proof
-    assume "a \<le> 0"
-    hence *: "Real a = 0" by simp
-    show ?thesis using `a \<le> 0` unfolding * by auto
-  next
-    assume "b \<le> 0"
-    hence *: "Real b = 0" by simp
-    show ?thesis using `b \<le> 0` unfolding * by auto
-  qed
-qed
-
-lemma real_pextreal_0[simp]: "real (0 :: pextreal) = 0"
-  unfolding zero_pextreal_def real_Real by simp
-
-lemma real_of_pextreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
-  by (cases X) auto
-
-lemma real_of_pextreal_eq: "real X = real Y \<longleftrightarrow>
-    (X = Y \<or> (X = 0 \<and> Y = \<omega>) \<or> (Y = 0 \<and> X = \<omega>))"
-  by (cases X, cases Y) (auto simp add: real_of_pextreal_eq_0)
-
-lemma real_of_pextreal_add: "real X + real Y =
-    (if X = \<omega> then real Y else if Y = \<omega> then real X else real (X + Y))"
-  by (auto simp: pextreal_noteq_omega_Ex)
-
-subsection "@{typ pextreal} is a monoid for multiplication"
-
-instantiation pextreal :: comm_monoid_mult
-begin
-
-definition "1 = Real 1"
-definition "x * y = (if x = 0 \<or> y = 0 then 0 else
-  pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
-
-lemma pextreal_times[simp]:
-  "Real r * Real p = (if 0 \<le> r \<and> 0 \<le> p then Real (r * p) else 0)"
-  "\<omega> * x = (if x = 0 then 0 else \<omega>)"
-  "x * \<omega> = (if x = 0 then 0 else \<omega>)"
-  "0 * x = 0"
-  "x * 0 = 0"
-  "1 = \<omega> \<longleftrightarrow> False"
-  "\<omega> = 1 \<longleftrightarrow> False"
-  by (auto simp add: times_pextreal_def one_pextreal_def)
-
-lemma pextreal_one_mult[simp]:
-  "Real x + 1 = (if 0 \<le> x then Real (x + 1) else 1)"
-  "1 + Real x = (if 0 \<le> x then Real (1 + x) else 1)"
-  unfolding one_pextreal_def by simp_all
-
-instance
-proof
-  fix a :: pextreal show "1 * a = a"
-    by (cases a) (simp_all add: one_pextreal_def)
-
-  fix b show "a * b = b * a"
-    by (cases a, cases b) (simp_all add: mult_nonneg_nonneg)
-
-  fix c show "a * b * c = a * (b * c)"
-    apply (cases a, cases b, cases c)
-    apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
-    apply (cases b, cases c)
-    apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
-    done
-qed
-end
-
-lemma pextreal_mult_cancel_left:
-  "a * b = a * c \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
-  by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
-
-lemma pextreal_mult_cancel_right:
-  "b * a = c * a \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
-  by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
-
-lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pextreal_def)
-
-lemma real_pextreal_1[simp]: "real (1 :: pextreal) = 1"
-  unfolding one_pextreal_def real_Real by simp
-
-lemma real_of_pextreal_mult: "real X * real Y = real (X * Y :: pextreal)"
-  by (cases X, cases Y) (auto simp: zero_le_mult_iff)
-
-lemma Real_mult_nonneg: assumes "x \<ge> 0" "y \<ge> 0"
-  shows "Real (x * y) = Real x * Real y" using assms by auto
-
-lemma Real_setprod: assumes "\<forall>x\<in>A. f x \<ge> 0" shows "Real (setprod f A) = setprod (\<lambda>x. Real (f x)) A"
-proof(cases "finite A")
-  case True thus ?thesis using assms
-  proof(induct A) case (insert x A)
-    have "0 \<le> setprod f A" apply(rule setprod_nonneg) using insert by auto
-    thus ?case unfolding setprod_insert[OF insert(1-2)] apply-
-      apply(subst Real_mult_nonneg) prefer 3 apply(subst insert(3)[THEN sym])
-      using insert by auto
-  qed auto
-qed auto
-
-subsection "@{typ pextreal} is a linear order"
-
-instantiation pextreal :: linorder
-begin
-
-definition "x < y \<longleftrightarrow> pextreal_case (\<lambda>i. pextreal_case (\<lambda>j. i < j) True y) False x"
-definition "x \<le> y \<longleftrightarrow> pextreal_case (\<lambda>j. pextreal_case (\<lambda>i. i \<le> j) False x) True y"
-
-lemma pextreal_less[simp]:
-  "Real r < \<omega>"
-  "Real r < Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r < p else 0 < p)"
-  "\<omega> < x \<longleftrightarrow> False"
-  "0 < \<omega>"
-  "0 < Real r \<longleftrightarrow> 0 < r"
-  "x < 0 \<longleftrightarrow> False"
-  "0 < (1::pextreal)"
-  by (simp_all add: less_pextreal_def zero_pextreal_def one_pextreal_def del: Real_0 Real_1)
-
-lemma pextreal_less_eq[simp]:
-  "x \<le> \<omega>"
-  "Real r \<le> Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r \<le> p else r \<le> 0)"
-  "0 \<le> x"
-  by (simp_all add: less_eq_pextreal_def zero_pextreal_def del: Real_0)
-
-lemma pextreal_\<omega>_less_eq[simp]:
-  "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
-  by (cases x) (simp_all add: not_le less_eq_pextreal_def)
-
-lemma pextreal_less_eq_zero[simp]:
-  "(x::pextreal) \<le> 0 \<longleftrightarrow> x = 0"
-  by (cases x) (simp_all add: zero_pextreal_def del: Real_0)
-
-instance
-proof
-  fix x :: pextreal
-  show "x \<le> x" by (cases x) simp_all
-  fix y
-  show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
-    by (cases x, cases y) auto
-  show "x \<le> y \<or> y \<le> x "
-    by (cases x, cases y) auto
-  { assume "x \<le> y" "y \<le> x" thus "x = y"
-      by (cases x, cases y) auto }
-  { fix z assume "x \<le> y" "y \<le> z"
-    thus "x \<le> z" by (cases x, cases y, cases z) auto }
-qed
-end
-
-lemma pextreal_zero_lessI[intro]:
-  "(a :: pextreal) \<noteq> 0 \<Longrightarrow> 0 < a"
-  by (cases a) auto
-
-lemma pextreal_less_omegaI[intro, simp]:
-  "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
-  by (cases a) auto
-
-lemma pextreal_plus_eq_0[simp]: "(a :: pextreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
-  by (cases a, cases b) auto
-
-lemma pextreal_le_add1[simp, intro]: "n \<le> n + (m::pextreal)"
-  by (cases n, cases m) simp_all
-
-lemma pextreal_le_add2: "(n::pextreal) + m \<le> k \<Longrightarrow> m \<le> k"
-  by (cases n, cases m, cases k) simp_all
-
-lemma pextreal_le_add3: "(n::pextreal) + m \<le> k \<Longrightarrow> n \<le> k"
-  by (cases n, cases m, cases k) simp_all
-
-lemma pextreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
-  by (cases x) auto
-
-lemma pextreal_0_less_mult_iff[simp]:
-  fixes x y :: pextreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
-  by (cases x, cases y) (auto simp: zero_less_mult_iff)
-
-lemma pextreal_ord_one[simp]:
-  "Real p < 1 \<longleftrightarrow> p < 1"
-  "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
-  "1 < Real p \<longleftrightarrow> 1 < p"
-  "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
-  by (simp_all add: one_pextreal_def del: Real_1)
-
-subsection {* @{text "x - y"} on @{typ pextreal} *}
-
-instantiation pextreal :: minus
-begin
-definition "x - y = (if y < x then THE d. x = y + d else 0 :: pextreal)"
-
-lemma minus_pextreal_eq:
-  "(x - y = (z :: pextreal)) \<longleftrightarrow> (if y < x then x = y + z else z = 0)"
-  (is "?diff \<longleftrightarrow> ?if")
-proof
-  assume ?diff
-  thus ?if
-  proof (cases "y < x")
-    case True
-    then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
-
-    show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pextreal_def
-    proof (rule theI2[where Q="\<lambda>d. x = y + d"])
-      show "x = y + pextreal_case (\<lambda>r. Real (r - real y)) \<omega> x" (is "x = y + ?d")
-        using `y < x` p by (cases x) simp_all
-
-      fix d assume "x = y + d"
-      thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all
-    qed simp
-  qed (simp add: minus_pextreal_def)
-next
-  assume ?if
-  thus ?diff
-  proof (cases "y < x")
-    case True
-    then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
-
-    from True `?if` have "x = y + z" by simp
-
-    show ?thesis unfolding minus_pextreal_def if_P[OF True] unfolding `x = y + z`
-    proof (rule the_equality)
-      fix d :: pextreal assume "y + z = y + d"
-      thus "d = z" using `y < x` p
-        by (cases d, cases z) simp_all
-    qed simp
-  qed (simp add: minus_pextreal_def)
-qed
-
-instance ..
-end
-
-lemma pextreal_minus[simp]:
-  "Real r - Real p = (if 0 \<le> r \<and> p < r then if 0 \<le> p then Real (r - p) else Real r else 0)"
-  "(A::pextreal) - A = 0"
-  "\<omega> - Real r = \<omega>"
-  "Real r - \<omega> = 0"
-  "A - 0 = A"
-  "0 - A = 0"
-  by (auto simp: minus_pextreal_eq not_less)
-
-lemma pextreal_le_epsilon:
-  fixes x y :: pextreal
-  assumes "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
-  shows "x \<le> y"
-proof (cases y)
-  case (preal r)
-  then obtain p where x: "x = Real p" "0 \<le> p"
-    using assms[of 1] by (cases x) auto
-  { fix e have "0 < e \<Longrightarrow> p \<le> r + e"
-      using assms[of "Real e"] preal x by auto }
-  hence "p \<le> r" by (rule field_le_epsilon)
-  thus ?thesis using preal x by auto
-qed simp
-
-instance pextreal :: "{ordered_comm_semiring, comm_semiring_1}"
-proof
-  show "0 \<noteq> (1::pextreal)" unfolding zero_pextreal_def one_pextreal_def
-    by (simp del: Real_1 Real_0)
-
-  fix a :: pextreal
-  show "0 * a = 0" "a * 0 = 0" by simp_all
-
-  fix b c :: pextreal
-  show "(a + b) * c = a * c + b * c"
-    by (cases c, cases a, cases b)
-       (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff)
-
-  { assume "a \<le> b" thus "c + a \<le> c + b"
-     by (cases c, cases a, cases b) auto }
-
-  assume "a \<le> b" "0 \<le> c"
-  thus "c * a \<le> c * b"
-    apply (cases c, cases a, cases b)
-    by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le)
-qed
-
-lemma mult_\<omega>[simp]: "x * y = \<omega> \<longleftrightarrow> (x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0"
-  by (cases x, cases y) auto
-
-lemma \<omega>_mult[simp]: "(\<omega> = x * y) = ((x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0)"
-  by (cases x, cases y) auto
-
-lemma pextreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pextreal) = 0"
-  by (cases x, cases y) (auto simp: mult_le_0_iff)
-
-lemma pextreal_mult_cancel:
-  fixes x y z :: pextreal
-  assumes "y \<le> z"
-  shows "x * y \<le> x * z"
-  using assms
-  by (cases x, cases y, cases z)
-     (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le)
-
-lemma Real_power[simp]:
-  "Real x ^ n = (if x \<le> 0 then (if n = 0 then 1 else 0) else Real (x ^ n))"
-  by (induct n) auto
-
-lemma Real_power_\<omega>[simp]:
-  "\<omega> ^ n = (if n = 0 then 1 else \<omega>)"
-  by (induct n) auto
-
-lemma pextreal_of_nat[simp]: "of_nat m = Real (real m)"
-  by (induct m) (auto simp: real_of_nat_Suc one_pextreal_def simp del: Real_1)
-
-lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
-proof safe
-  assume "x < \<omega>"
-  then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
-  moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
-  ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
-qed auto
-
-lemma real_of_pextreal_mono:
-  fixes a b :: pextreal
-  assumes "b \<noteq> \<omega>" "a \<le> b"
-  shows "real a \<le> real b"
-using assms by (cases b, cases a) auto
-
-lemma setprod_pextreal_0:
-  "(\<Prod>i\<in>I. f i) = (0::pextreal) \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = 0)"
-proof cases
-  assume "finite I" then show ?thesis
-  proof (induct I)
-    case (insert i I)
-    then show ?case by simp
-  qed simp
-qed simp
-
-lemma setprod_\<omega>:
-  "(\<Prod>i\<in>I. f i) = \<omega> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<omega>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
-proof cases
-  assume "finite I" then show ?thesis
-  proof (induct I)
-    case (insert i I) then show ?case
-      by (auto simp: setprod_pextreal_0)
-  qed simp
-qed simp
-
-instance pextreal :: "semiring_char_0"
-proof
-  fix m n
-  show "inj (of_nat::nat\<Rightarrow>pextreal)" by (auto intro!: inj_onI)
-qed
-
-subsection "@{typ pextreal} is a complete lattice"
-
-instantiation pextreal :: lattice
-begin
-definition [simp]: "sup x y = (max x y :: pextreal)"
-definition [simp]: "inf x y = (min x y :: pextreal)"
-instance proof qed simp_all
-end
-
-instantiation pextreal :: complete_lattice
-begin
-
-definition "bot = Real 0"
-definition "top = \<omega>"
-
-definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pextreal)"
-definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pextreal)"
-
-lemma pextreal_complete_Sup:
-  fixes S :: "pextreal set" assumes "S \<noteq> {}"
-  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
-proof (cases "\<exists>x\<ge>0. \<forall>a\<in>S. a \<le> Real x")
-  case False
-  hence *: "\<And>x. x\<ge>0 \<Longrightarrow> \<exists>a\<in>S. \<not>a \<le> Real x" by simp
-  show ?thesis
-  proof (safe intro!: exI[of _ \<omega>])
-    fix y assume **: "\<forall>z\<in>S. z \<le> y"
-    show "\<omega> \<le> y" unfolding pextreal_\<omega>_less_eq
-    proof (rule ccontr)
-      assume "y \<noteq> \<omega>"
-      then obtain x where [simp]: "y = Real x" and "0 \<le> x" by (cases y) auto
-      from *[OF `0 \<le> x`] show False using ** by auto
-    qed
-  qed simp
-next
-  case True then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> Real y" and "0 \<le> y" by auto
-  from y[of \<omega>] have "\<omega> \<notin> S" by auto
-
-  with `S \<noteq> {}` obtain x where "x \<in> S" and "x \<noteq> \<omega>" by auto
-
-  have bound: "\<forall>x\<in>real ` S. x \<le> y"
-  proof
-    fix z assume "z \<in> real ` S" then guess a ..
-    with y[of a] `\<omega> \<notin> S` `0 \<le> y` show "z \<le> y" by (cases a) auto
-  qed
-  with reals_complete2[of "real ` S"] `x \<in> S`
-  obtain s where s: "\<forall>y\<in>S. real y \<le> s" "\<forall>z. ((\<forall>y\<in>S. real y \<le> z) \<longrightarrow> s \<le> z)"
-    by auto
-
-  show ?thesis
-  proof (safe intro!: exI[of _ "Real s"])
-    fix z assume "z \<in> S" thus "z \<le> Real s"
-      using s `\<omega> \<notin> S` by (cases z) auto
-  next
-    fix z assume *: "\<forall>y\<in>S. y \<le> z"
-    show "Real s \<le> z"
-    proof (cases z)
-      case (preal u)
-      { fix v assume "v \<in> S"
-        hence "v \<le> Real u" using * preal by auto
-        hence "real v \<le> u" using `\<omega> \<notin> S` `0 \<le> u` by (cases v) auto }
-      hence "s \<le> u" using s(2) by auto
-      thus "Real s \<le> z" using preal by simp
-    qed simp
-  qed
-qed
-
-lemma pextreal_complete_Inf:
-  fixes S :: "pextreal set" assumes "S \<noteq> {}"
-  shows "\<exists>x. (\<forall>y\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> x)"
-proof (cases "S = {\<omega>}")
-  case True thus ?thesis by (auto intro!: exI[of _ \<omega>])
-next
-  case False with `S \<noteq> {}` have "S - {\<omega>} \<noteq> {}" by auto
-  hence not_empty: "\<exists>x. x \<in> uminus ` real ` (S - {\<omega>})" by auto
-  have bounds: "\<exists>x. \<forall>y\<in>uminus ` real ` (S - {\<omega>}). y \<le> x" by (auto intro!: exI[of _ 0])
-  from reals_complete2[OF not_empty bounds]
-  obtain s where s: "\<And>y. y\<in>S - {\<omega>} \<Longrightarrow> - real y \<le> s" "\<forall>z. ((\<forall>y\<in>S - {\<omega>}. - real y \<le> z) \<longrightarrow> s \<le> z)"
-    by auto
-
-  show ?thesis
-  proof (safe intro!: exI[of _ "Real (-s)"])
-    fix z assume "z \<in> S"
-    show "Real (-s) \<le> z"
-    proof (cases z)
-      case (preal r)
-      with s `z \<in> S` have "z \<in> S - {\<omega>}" by simp
-      hence "- r \<le> s" using preal s(1)[of z] by auto
-      hence "- s \<le> r" by (subst neg_le_iff_le[symmetric]) simp
-      thus ?thesis using preal by simp
-    qed simp
-  next
-    fix z assume *: "\<forall>y\<in>S. z \<le> y"
-    show "z \<le> Real (-s)"
-    proof (cases z)
-      case (preal u)
-      { fix v assume "v \<in> S-{\<omega>}"
-        hence "Real u \<le> v" using * preal by auto
-        hence "- real v \<le> - u" using `0 \<le> u` `v \<in> S - {\<omega>}` by (cases v) auto }
-      hence "u \<le> - s" using s(2) by (subst neg_le_iff_le[symmetric]) auto
-      thus "z \<le> Real (-s)" using preal by simp
-    next
-      case infinite
-      with * have "S = {\<omega>}" using `S \<noteq> {}` by auto
-      with `S - {\<omega>} \<noteq> {}` show ?thesis by simp
-    qed
-  qed
-qed
-
-instance
-proof
-  fix x :: pextreal and A
-
-  show "bot \<le> x" by (cases x) (simp_all add: bot_pextreal_def)
-  show "x \<le> top" by (simp add: top_pextreal_def)
-
-  { assume "x \<in> A"
-    with pextreal_complete_Sup[of A]
-    obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
-    hence "x \<le> s" using `x \<in> A` by auto
-    also have "... = Sup A" using s unfolding Sup_pextreal_def
-      by (auto intro!: Least_equality[symmetric])
-    finally show "x \<le> Sup A" . }
-
-  { assume "x \<in> A"
-    with pextreal_complete_Inf[of A]
-    obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
-    hence "Inf A = i" unfolding Inf_pextreal_def
-      by (auto intro!: Greatest_equality)
-    also have "i \<le> x" using i `x \<in> A` by auto
-    finally show "Inf A \<le> x" . }
-
-  { assume *: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x"
-    show "Sup A \<le> x"
-    proof (cases "A = {}")
-      case True
-      hence "Sup A = 0" unfolding Sup_pextreal_def
-        by (auto intro!: Least_equality)
-      thus "Sup A \<le> x" by simp
-    next
-      case False
-      with pextreal_complete_Sup[of A]
-      obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
-      hence "Sup A = s"
-        unfolding Sup_pextreal_def by (auto intro!: Least_equality)
-      also have "s \<le> x" using * s by auto
-      finally show "Sup A \<le> x" .
-    qed }
-
-  { assume *: "\<And>z. z \<in> A \<Longrightarrow> x \<le> z"
-    show "x \<le> Inf A"
-    proof (cases "A = {}")
-      case True
-      hence "Inf A = \<omega>" unfolding Inf_pextreal_def
-        by (auto intro!: Greatest_equality)
-      thus "x \<le> Inf A" by simp
-    next
-      case False
-      with pextreal_complete_Inf[of A]
-      obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
-      have "x \<le> i" using * i by auto
-      also have "i = Inf A" using i
-        unfolding Inf_pextreal_def by (auto intro!: Greatest_equality[symmetric])
-      finally show "x \<le> Inf A" .
-    qed }
-qed
-end
-
-lemma Inf_pextreal_iff:
-  fixes z :: pextreal
-  shows "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> (\<exists>x\<in>X. x<y) \<longleftrightarrow> Inf X < y"
-  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
-            order_less_le_trans)
-
-lemma Inf_greater:
-  fixes z :: pextreal assumes "Inf X < z"
-  shows "\<exists>x \<in> X. x < z"
-proof -
-  have "X \<noteq> {}" using assms by (auto simp: Inf_empty top_pextreal_def)
-  with assms show ?thesis
-    by (metis Inf_pextreal_iff mem_def not_leE)
-qed
-
-lemma Inf_close:
-  fixes e :: pextreal assumes "Inf X \<noteq> \<omega>" "0 < e"
-  shows "\<exists>x \<in> X. x < Inf X + e"
-proof (rule Inf_greater)
-  show "Inf X < Inf X + e" using assms
-    by (cases "Inf X", cases e) auto
-qed
-
-lemma pextreal_SUPI:
-  fixes x :: pextreal
-  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
-  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y"
-  shows "(SUP i:A. f i) = x"
-  unfolding SUPR_def Sup_pextreal_def
-  using assms by (auto intro!: Least_equality)
-
-lemma Sup_pextreal_iff:
-  fixes z :: pextreal
-  shows "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> (\<exists>x\<in>X. y<x) \<longleftrightarrow> y < Sup X"
-  by (metis complete_lattice_class.Sup_least complete_lattice_class.Sup_upper less_le_not_le linear
-            order_less_le_trans)
-
-lemma Sup_lesser:
-  fixes z :: pextreal assumes "z < Sup X"
-  shows "\<exists>x \<in> X. z < x"
-proof -
-  have "X \<noteq> {}" using assms by (auto simp: Sup_empty bot_pextreal_def)
-  with assms show ?thesis
-    by (metis Sup_pextreal_iff mem_def not_leE)
-qed
-
-lemma Sup_eq_\<omega>: "\<omega> \<in> S \<Longrightarrow> Sup S = \<omega>"
-  unfolding Sup_pextreal_def
-  by (auto intro!: Least_equality)
-
-lemma Sup_close:
-  assumes "0 < e" and S: "Sup S \<noteq> \<omega>" "S \<noteq> {}"
-  shows "\<exists>X\<in>S. Sup S < X + e"
-proof cases
-  assume "Sup S = 0"
-  moreover obtain X where "X \<in> S" using `S \<noteq> {}` by auto
-  ultimately show ?thesis using `0 < e` by (auto intro!: bexI[OF _ `X\<in>S`])
-next
-  assume "Sup S \<noteq> 0"
-  have "\<exists>X\<in>S. Sup S - e < X"
-  proof (rule Sup_lesser)
-    show "Sup S - e < Sup S" using `0 < e` `Sup S \<noteq> 0` `Sup S \<noteq> \<omega>`
-      by (cases e) (auto simp: pextreal_noteq_omega_Ex)
-  qed
-  then guess X .. note X = this
-  with `Sup S \<noteq> \<omega>` Sup_eq_\<omega> have "X \<noteq> \<omega>" by auto
-  thus ?thesis using `Sup S \<noteq> \<omega>` X unfolding pextreal_noteq_omega_Ex
-    by (cases e) (auto intro!: bexI[OF _ `X\<in>S`] simp: split: split_if_asm)
-qed
-
-lemma Sup_\<omega>: "(SUP i::nat. Real (real i)) = \<omega>"
-proof (rule pextreal_SUPI)
-  fix y assume *: "\<And>i::nat. i \<in> UNIV \<Longrightarrow> Real (real i) \<le> y"
-  thus "\<omega> \<le> y"
-  proof (cases y)
-    case (preal r)
-    then obtain k :: nat where "r < real k"
-      using ex_less_of_nat by (auto simp: real_eq_of_nat)
-    with *[of k] preal show ?thesis by auto
-  qed simp
-qed simp
-
-lemma SUP_\<omega>: "(SUP i:A. f i) = \<omega> \<longleftrightarrow> (\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)"
-proof
-  assume *: "(SUP i:A. f i) = \<omega>"
-  show "(\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)" unfolding *[symmetric]
-  proof (intro allI impI)
-    fix x assume "x < SUPR A f" then show "\<exists>i\<in>A. x < f i"
-      unfolding less_SUP_iff by auto
-  qed
-next
-  assume *: "\<forall>x<\<omega>. \<exists>i\<in>A. x < f i"
-  show "(SUP i:A. f i) = \<omega>"
-  proof (rule pextreal_SUPI)
-    fix y assume **: "\<And>i. i \<in> A \<Longrightarrow> f i \<le> y"
-    show "\<omega> \<le> y"
-    proof cases
-      assume "y < \<omega>"
-      from *[THEN spec, THEN mp, OF this]
-      obtain i where "i \<in> A" "\<not> (f i \<le> y)" by auto
-      with ** show ?thesis by auto
-    qed auto
-  qed auto
-qed
-
-subsubsection {* Equivalence between @{text "f ----> x"} and @{text SUP} on @{typ pextreal} *}
-
-lemma monoseq_monoI: "mono f \<Longrightarrow> monoseq f"
-  unfolding mono_def monoseq_def by auto
-
-lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
-  unfolding mono_def incseq_def by auto
-
-lemma SUP_eq_LIMSEQ:
-  assumes "mono f" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
-  shows "(SUP n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
-proof
-  assume x: "(SUP n. Real (f n)) = Real x"
-  { fix n
-    have "Real (f n) \<le> Real x" using x[symmetric] by (auto intro: le_SUPI)
-    hence "f n \<le> x" using assms by simp }
-  show "f ----> x"
-  proof (rule LIMSEQ_I)
-    fix r :: real assume "0 < r"
-    show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
-    proof (rule ccontr)
-      assume *: "\<not> ?thesis"
-      { fix N
-        from * obtain n where "N \<le> n" "r \<le> x - f n"
-          using `\<And>n. f n \<le> x` by (auto simp: not_less)
-        hence "f N \<le> f n" using `mono f` by (auto dest: monoD)
-        hence "f N \<le> x - r" using `r \<le> x - f n` by auto
-        hence "Real (f N) \<le> Real (x - r)" and "r \<le> x" using `0 \<le> f N` by auto }
-      hence "(SUP n. Real (f n)) \<le> Real (x - r)"
-        and "Real (x - r) < Real x" using `0 < r` by (auto intro: SUP_leI)
-      hence "(SUP n. Real (f n)) < Real x" by (rule le_less_trans)
-      thus False using x by auto
-    qed
-  qed
-next
-  assume "f ----> x"
-  show "(SUP n. Real (f n)) = Real x"
-  proof (rule pextreal_SUPI)
-    fix n
-    from incseq_le[of f x] `mono f` `f ----> x`
-    show "Real (f n) \<le> Real x" using assms incseq_mono by auto
-  next
-    fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> Real (f n) \<le> y"
-    show "Real x \<le> y"
-    proof (cases y)
-      case (preal r)
-      with * have "\<exists>N. \<forall>n\<ge>N. f n \<le> r" using assms by fastsimp
-      from LIMSEQ_le_const2[OF `f ----> x` this]
-      show "Real x \<le> y" using `0 \<le> x` preal by auto
-    qed simp
-  qed
-qed
-
-lemma SUPR_bound:
-  assumes "\<forall>N. f N \<le> x"
-  shows "(SUP n. f n) \<le> x"
-  using assms by (simp add: SUPR_def Sup_le_iff)
-
-lemma pextreal_less_eq_diff_eq_sum:
-  fixes x y z :: pextreal
-  assumes "y \<le> x" and "x \<noteq> \<omega>"
-  shows "z \<le> x - y \<longleftrightarrow> z + y \<le> x"
-  using assms
-  apply (cases z, cases y, cases x)
-  by (simp_all add: field_simps minus_pextreal_eq)
-
-lemma Real_diff_less_omega: "Real r - x < \<omega>" by (cases x) auto
-
-subsubsection {* Numbers on @{typ pextreal} *}
-
-instantiation pextreal :: number
-begin
-definition [simp]: "number_of x = Real (number_of x)"
-instance proof qed
-end
-
-subsubsection {* Division on @{typ pextreal} *}
-
-instantiation pextreal :: inverse
-begin
-
-definition "inverse x = pextreal_case (\<lambda>x. if x = 0 then \<omega> else Real (inverse x)) 0 x"
-definition [simp]: "x / y = x * inverse (y :: pextreal)"
-
-instance proof qed
-end
-
-lemma pextreal_inverse[simp]:
-  "inverse 0 = \<omega>"
-  "inverse (Real x) = (if x \<le> 0 then \<omega> else Real (inverse x))"
-  "inverse \<omega> = 0"
-  "inverse (1::pextreal) = 1"
-  "inverse (inverse x) = x"
-  by (simp_all add: inverse_pextreal_def one_pextreal_def split: pextreal_case_split del: Real_1)
-
-lemma pextreal_inverse_le_eq:
-  assumes "x \<noteq> 0" "x \<noteq> \<omega>"
-  shows "y \<le> z / x \<longleftrightarrow> x * y \<le> (z :: pextreal)"
-proof -
-  from assms obtain r where r: "x = Real r" "0 < r" by (cases x) auto
-  { fix p q :: real assume "0 \<le> p" "0 \<le> q"
-    have "p \<le> q * inverse r \<longleftrightarrow> p \<le> q / r" by (simp add: divide_inverse)
-    also have "... \<longleftrightarrow> p * r \<le> q" using `0 < r` by (auto simp: field_simps)
-    finally have "p \<le> q * inverse r \<longleftrightarrow> p * r \<le> q" . }
-  with r show ?thesis
-    by (cases y, cases z, auto simp: zero_le_mult_iff field_simps)
-qed
-
-lemma inverse_antimono_strict:
-  fixes x y :: pextreal
-  assumes "x < y" shows "inverse y < inverse x"
-  using assms by (cases x, cases y) auto
-
-lemma inverse_antimono:
-  fixes x y :: pextreal
-  assumes "x \<le> y" shows "inverse y \<le> inverse x"
-  using assms by (cases x, cases y) auto
-
-lemma pextreal_inverse_\<omega>_iff[simp]: "inverse x = \<omega> \<longleftrightarrow> x = 0"
-  by (cases x) auto
-
-subsection "Infinite sum over @{typ pextreal}"
-
-text {*
-
-The infinite sum over @{typ pextreal} has the nice property that it is always
-defined.
-
-*}
-
-definition psuminf :: "(nat \<Rightarrow> pextreal) \<Rightarrow> pextreal" (binder "\<Sum>\<^isub>\<infinity>" 10) where
-  "(\<Sum>\<^isub>\<infinity> x. f x) = (SUP n. \<Sum>i<n. f i)"
-
-subsubsection {* Equivalence between @{text "\<Sum> n. f n"} and @{text "\<Sum>\<^isub>\<infinity> n. f n"} *}
-
-lemma setsum_Real:
-  assumes "\<forall>x\<in>A. 0 \<le> f x"
-  shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
-proof (cases "finite A")
-  case True
-  thus ?thesis using assms
-  proof induct case (insert x s)
-    hence "0 \<le> setsum f s" apply-apply(rule setsum_nonneg) by auto
-    thus ?case using insert by auto
-  qed auto
-qed simp
-
-lemma setsum_Real':
-  assumes "\<forall>x. 0 \<le> f x"
-  shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
-  apply(rule setsum_Real) using assms by auto
-
-lemma setsum_\<omega>:
-  "(\<Sum>x\<in>P. f x) = \<omega> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<omega>))"
-proof safe
-  assume *: "setsum f P = \<omega>"
-  show "finite P"
-  proof (rule ccontr) assume "infinite P" with * show False by auto qed
-  show "\<exists>i\<in>P. f i = \<omega>"
-  proof (rule ccontr)
-    assume "\<not> ?thesis" hence "\<And>i. i\<in>P \<Longrightarrow> f i \<noteq> \<omega>" by auto
-    from `finite P` this have "setsum f P \<noteq> \<omega>"
-      by induct auto
-    with * show False by auto
-  qed
-next
-  fix i assume "finite P" "i \<in> P" "f i = \<omega>"
-  thus "setsum f P = \<omega>"
-  proof induct
-    case (insert x A)
-    show ?case using insert by (cases "x = i") auto
-  qed simp
-qed
-
-lemma real_of_pextreal_setsum:
-  assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> \<omega>"
-  shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
-proof cases
-  assume "finite S"
-  from this assms show ?thesis
-    by induct (simp_all add: real_of_pextreal_add setsum_\<omega>)
-qed simp
-
-lemma setsum_0:
-  fixes f :: "'a \<Rightarrow> pextreal" assumes "finite A"
-  shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
-  using assms by induct auto
-
-lemma suminf_imp_psuminf:
-  assumes "f sums x" and "\<forall>n. 0 \<le> f n"
-  shows "(\<Sum>\<^isub>\<infinity> x. Real (f x)) = Real x"
-  unfolding psuminf_def setsum_Real'[OF assms(2)]
-proof (rule SUP_eq_LIMSEQ[THEN iffD2])
-  show "mono (\<lambda>n. setsum f {..<n})" (is "mono ?S")
-    unfolding mono_iff_le_Suc using assms by simp
-
-  { fix n show "0 \<le> ?S n"
-      using setsum_nonneg[of "{..<n}" f] assms by auto }
-
-  thus "0 \<le> x" "?S ----> x"
-    using `f sums x` LIMSEQ_le_const
-    by (auto simp: atLeast0LessThan sums_def)
-qed
-
-lemma psuminf_equality:
-  assumes "\<And>n. setsum f {..<n} \<le> x"
-  and "\<And>y. y \<noteq> \<omega> \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> y) \<Longrightarrow> x \<le> y"
-  shows "psuminf f = x"
-  unfolding psuminf_def
-proof (safe intro!: pextreal_SUPI)
-  fix n show "setsum f {..<n} \<le> x" using assms(1) .
-next
-  fix y assume *: "\<forall>n. n \<in> UNIV \<longrightarrow> setsum f {..<n} \<le> y"
-  show "x \<le> y"
-  proof (cases "y = \<omega>")
-    assume "y \<noteq> \<omega>" from assms(2)[OF this] *
-    show "x \<le> y" by auto
-  qed simp
-qed
-
-lemma psuminf_\<omega>:
-  assumes "f i = \<omega>"
-  shows "(\<Sum>\<^isub>\<infinity> x. f x) = \<omega>"
-proof (rule psuminf_equality)
-  fix y assume *: "\<And>n. setsum f {..<n} \<le> y"
-  have "setsum f {..<Suc i} = \<omega>" 
-    using assms by (simp add: setsum_\<omega>)
-  thus "\<omega> \<le> y" using *[of "Suc i"] by auto
-qed simp
-
-lemma psuminf_imp_suminf:
-  assumes "(\<Sum>\<^isub>\<infinity> x. f x) \<noteq> \<omega>"
-  shows "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity> x. f x)"
-proof -
-  have "\<forall>i. \<exists>r. f i = Real r \<and> 0 \<le> r"
-  proof
-    fix i show "\<exists>r. f i = Real r \<and> 0 \<le> r" using psuminf_\<omega> assms by (cases "f i") auto
-  qed
-  from choice[OF this] obtain r where f: "f = (\<lambda>i. Real (r i))"
-    and pos: "\<forall>i. 0 \<le> r i"
-    by (auto simp: fun_eq_iff)
-  hence [simp]: "\<And>i. real (f i) = r i" by auto
-
-  have "mono (\<lambda>n. setsum r {..<n})" (is "mono ?S")
-    unfolding mono_iff_le_Suc using pos by simp
-
-  { fix n have "0 \<le> ?S n"
-      using setsum_nonneg[of "{..<n}" r] pos by auto }
-
-  from assms obtain p where *: "(\<Sum>\<^isub>\<infinity> x. f x) = Real p" and "0 \<le> p"
-    by (cases "(\<Sum>\<^isub>\<infinity> x. f x)") auto
-  show ?thesis unfolding * using * pos `0 \<le> p` SUP_eq_LIMSEQ[OF `mono ?S` `\<And>n. 0 \<le> ?S n` `0 \<le> p`]
-    by (simp add: f atLeast0LessThan sums_def psuminf_def setsum_Real'[OF pos] f)
-qed
-
-lemma psuminf_bound:
-  assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x"
-  shows "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x"
-  using assms by (simp add: psuminf_def SUPR_def Sup_le_iff)
-
-lemma psuminf_bound_add:
-  assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x"
-  shows "(\<Sum>\<^isub>\<infinity> n. f n) + y \<le> x"
-proof (cases "x = \<omega>")
-  have "y \<le> x" using assms by (auto intro: pextreal_le_add2)
-  assume "x \<noteq> \<omega>"
-  note move_y = pextreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
-
-  have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" using assms by (simp add: move_y)
-  hence "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x - y" by (rule psuminf_bound)
-  thus ?thesis by (simp add: move_y)
-qed simp
-
-lemma psuminf_finite:
-  assumes "\<forall>N\<ge>n. f N = 0"
-  shows "(\<Sum>\<^isub>\<infinity> n. f n) = (\<Sum>N<n. f N)"
-proof (rule psuminf_equality)
-  fix N
-  show "setsum f {..<N} \<le> setsum f {..<n}"
-  proof (cases rule: linorder_cases)
-    assume "N < n" thus ?thesis by (auto intro!: setsum_mono3)
-  next
-    assume "n < N"
-    hence *: "{..<N} = {..<n} \<union> {n..<N}" by auto
-    moreover have "setsum f {n..<N} = 0"
-      using assms by (auto intro!: setsum_0')
-    ultimately show ?thesis unfolding *
-      by (subst setsum_Un_disjoint) auto
-  qed simp
-qed simp
-
-lemma psuminf_upper:
-  shows "(\<Sum>n<N. f n) \<le> (\<Sum>\<^isub>\<infinity> n. f n)"
-  unfolding psuminf_def SUPR_def
-  by (auto intro: complete_lattice_class.Sup_upper image_eqI)
-
-lemma psuminf_le:
-  assumes "\<And>N. f N \<le> g N"
-  shows "psuminf f \<le> psuminf g"
-proof (safe intro!: psuminf_bound)
-  fix n
-  have "setsum f {..<n} \<le> setsum g {..<n}" using assms by (auto intro: setsum_mono)
-  also have "... \<le> psuminf g" by (rule psuminf_upper)
-  finally show "setsum f {..<n} \<le> psuminf g" .
-qed
-
-lemma psuminf_const[simp]: "psuminf (\<lambda>n. c) = (if c = 0 then 0 else \<omega>)" (is "_ = ?if")
-proof (rule psuminf_equality)
-  fix y assume *: "\<And>n :: nat. (\<Sum>n<n. c) \<le> y" and "y \<noteq> \<omega>"
-  then obtain r p where
-    y: "y = Real r" "0 \<le> r" and
-    c: "c = Real p" "0 \<le> p"
-      using *[of 1] by (cases c, cases y) auto
-  show "(if c = 0 then 0 else \<omega>) \<le> y"
-  proof (cases "p = 0")
-    assume "p = 0" with c show ?thesis by simp
-  next
-    assume "p \<noteq> 0"
-    with * c y have **: "\<And>n :: nat. real n \<le> r / p"
-      by (auto simp: zero_le_mult_iff field_simps)
-    from ex_less_of_nat[of "r / p"] guess n ..
-    with **[of n] show ?thesis by (simp add: real_eq_of_nat)
-  qed
-qed (cases "c = 0", simp_all)
-
-lemma psuminf_add[simp]: "psuminf (\<lambda>n. f n + g n) = psuminf f + psuminf g"
-proof (rule psuminf_equality)
-  fix n
-  from psuminf_upper[of f n] psuminf_upper[of g n]
-  show "(\<Sum>n<n. f n + g n) \<le> psuminf f + psuminf g"
-    by (auto simp add: setsum_addf intro!: add_mono)
-next
-  fix y assume *: "\<And>n. (\<Sum>n<n. f n + g n) \<le> y" and "y \<noteq> \<omega>"
-  { fix n m
-    have **: "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y"
-    proof (cases rule: linorder_le_cases)
-      assume "n \<le> m"
-      hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<m. f n) + (\<Sum>n<m. g n)"
-        by (auto intro!: add_right_mono setsum_mono3)
-      also have "... \<le> y"
-        using * by (simp add: setsum_addf)
-      finally show ?thesis .
-    next
-      assume "m \<le> n"
-      hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<n. f n) + (\<Sum>n<n. g n)"
-        by (auto intro!: add_left_mono setsum_mono3)
-      also have "... \<le> y"
-        using * by (simp add: setsum_addf)
-      finally show ?thesis .
-    qed }
-  hence "\<And>m. \<forall>n. (\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y" by simp
-  from psuminf_bound_add[OF this]
-  have "\<forall>m. (\<Sum>n<m. g n) + psuminf f \<le> y" by (simp add: ac_simps)
-  from psuminf_bound_add[OF this]
-  show "psuminf f + psuminf g \<le> y" by (simp add: ac_simps)
-qed
-
-lemma psuminf_0: "psuminf f = 0 \<longleftrightarrow> (\<forall>i. f i = 0)"
-proof safe
-  assume "\<forall>i. f i = 0"
-  hence "f = (\<lambda>i. 0)" by (simp add: fun_eq_iff)
-  thus "psuminf f = 0" using psuminf_const by simp
-next
-  fix i assume "psuminf f = 0"
-  hence "(\<Sum>n<Suc i. f n) = 0" using psuminf_upper[of f "Suc i"] by simp
-  thus "f i = 0" by simp
-qed
-
-lemma psuminf_cmult_right[simp]: "psuminf (\<lambda>n. c * f n) = c * psuminf f"
-proof (rule psuminf_equality)
-  fix n show "(\<Sum>n<n. c * f n) \<le> c * psuminf f"
-   by (auto simp: setsum_right_distrib[symmetric] intro: mult_left_mono psuminf_upper)
-next
-  fix y
-  assume "\<And>n. (\<Sum>n<n. c * f n) \<le> y"
-  hence *: "\<And>n. c * (\<Sum>n<n. f n) \<le> y" by (auto simp add: setsum_right_distrib)
-  thus "c * psuminf f \<le> y"
-  proof (cases "c = \<omega> \<or> c = 0")
-    assume "c = \<omega> \<or> c = 0"
-    thus ?thesis
-      using * by (fastsimp simp add: psuminf_0 setsum_0 split: split_if_asm)
-  next
-    assume "\<not> (c = \<omega> \<or> c = 0)"
-    hence "c \<noteq> 0" "c \<noteq> \<omega>" by auto
-    note rewrite_div = pextreal_inverse_le_eq[OF this, of _ y]
-    hence "\<forall>n. (\<Sum>n<n. f n) \<le> y / c" using * by simp
-    hence "psuminf f \<le> y / c" by (rule psuminf_bound)
-    thus ?thesis using rewrite_div by simp
-  qed
-qed
-
-lemma psuminf_cmult_left[simp]: "psuminf (\<lambda>n. f n * c) = psuminf f * c"
-  using psuminf_cmult_right[of c f] by (simp add: ac_simps)
-
-lemma psuminf_half_series: "psuminf (\<lambda>n. (1/2)^Suc n) = 1"
-  using suminf_imp_psuminf[OF power_half_series] by auto
-
-lemma setsum_pinfsum: "(\<Sum>\<^isub>\<infinity> n. \<Sum>m\<in>A. f n m) = (\<Sum>m\<in>A. (\<Sum>\<^isub>\<infinity> n. f n m))"
-proof (cases "finite A")
-  assume "finite A"
-  thus ?thesis by induct simp_all
-qed simp
-
-lemma psuminf_reindex:
-  fixes f:: "nat \<Rightarrow> nat" assumes "bij f"
-  shows "psuminf (g \<circ> f) = psuminf g"
-proof -
-  have [intro, simp]: "\<And>A. inj_on f A" using `bij f` unfolding bij_def by (auto intro: subset_inj_on)
-  have f[intro, simp]: "\<And>x. f (inv f x) = x"
-    using `bij f` unfolding bij_def by (auto intro: surj_f_inv_f)
-  show ?thesis
-  proof (rule psuminf_equality)
-    fix n
-    have "setsum (g \<circ> f) {..<n} = setsum g (f ` {..<n})"
-      by (simp add: setsum_reindex)
-    also have "\<dots> \<le> setsum g {..Max (f ` {..<n})}"
-      by (rule setsum_mono3) auto
-    also have "\<dots> \<le> psuminf g" unfolding lessThan_Suc_atMost[symmetric] by (rule psuminf_upper)
-    finally show "setsum (g \<circ> f) {..<n} \<le> psuminf g" .
-  next
-    fix y assume *: "\<And>n. setsum (g \<circ> f) {..<n} \<le> y"
-    show "psuminf g \<le> y"
-    proof (safe intro!: psuminf_bound)
-      fix N
-      have "setsum g {..<N} \<le> setsum g (f ` {..Max (inv f ` {..<N})})"
-        by (rule setsum_mono3) (auto intro!: image_eqI[where f="f", OF f[symmetric]])
-      also have "\<dots> = setsum (g \<circ> f) {..Max (inv f ` {..<N})}"
-        by (simp add: setsum_reindex)
-      also have "\<dots> \<le> y" unfolding lessThan_Suc_atMost[symmetric] by (rule *)
-      finally show "setsum g {..<N} \<le> y" .
-    qed
-  qed
-qed
-
-lemma pextreal_mult_less_right:
-  assumes "b * a < c * a" "0 < a" "a < \<omega>"
-  shows "b < c"
-  using assms
-  by (cases a, cases b, cases c) (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
-
-lemma pextreal_\<omega>_eq_plus[simp]: "\<omega> = a + b \<longleftrightarrow> (a = \<omega> \<or> b = \<omega>)"
-  by (cases a, cases b) auto
-
-lemma pextreal_of_nat_le_iff:
-  "(of_nat k :: pextreal) \<le> of_nat m \<longleftrightarrow> k \<le> m" by auto
-
-lemma pextreal_of_nat_less_iff:
-  "(of_nat k :: pextreal) < of_nat m \<longleftrightarrow> k < m" by auto
-
-lemma pextreal_bound_add:
-  assumes "\<forall>N. f N + y \<le> (x::pextreal)"
-  shows "(SUP n. f n) + y \<le> x"
-proof (cases "x = \<omega>")
-  have "y \<le> x" using assms by (auto intro: pextreal_le_add2)
-  assume "x \<noteq> \<omega>"
-  note move_y = pextreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
-
-  have "\<forall>N. f N \<le> x - y" using assms by (simp add: move_y)
-  hence "(SUP n. f n) \<le> x - y" by (rule SUPR_bound)
-  thus ?thesis by (simp add: move_y)
-qed simp
-
-lemma SUPR_pextreal_add:
-  fixes f g :: "nat \<Rightarrow> pextreal"
-  assumes f: "\<forall>n. f n \<le> f (Suc n)" and g: "\<forall>n. g n \<le> g (Suc n)"
-  shows "(SUP n. f n + g n) = (SUP n. f n) + (SUP n. g n)"
-proof (rule pextreal_SUPI)
-  fix n :: nat from le_SUPI[of n UNIV f] le_SUPI[of n UNIV g]
-  show "f n + g n \<le> (SUP n. f n) + (SUP n. g n)"
-    by (auto intro!: add_mono)
-next
-  fix y assume *: "\<And>n. n \<in> UNIV \<Longrightarrow> f n + g n \<le> y"
-  { fix n m
-    have "f n + g m \<le> y"
-    proof (cases rule: linorder_le_cases)
-      assume "n \<le> m"
-      hence "f n + g m \<le> f m + g m"
-        using f lift_Suc_mono_le by (auto intro!: add_right_mono)
-      also have "\<dots> \<le> y" using * by simp
-      finally show ?thesis .
-    next
-      assume "m \<le> n"
-      hence "f n + g m \<le> f n + g n"
-        using g lift_Suc_mono_le by (auto intro!: add_left_mono)
-      also have "\<dots> \<le> y" using * by simp
-      finally show ?thesis .
-    qed }
-  hence "\<And>m. \<forall>n. f n + g m \<le> y" by simp
-  from pextreal_bound_add[OF this]
-  have "\<forall>m. (g m) + (SUP n. f n) \<le> y" by (simp add: ac_simps)
-  from pextreal_bound_add[OF this]
-  show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
-qed
-
-lemma SUPR_pextreal_setsum:
-  fixes f :: "'x \<Rightarrow> nat \<Rightarrow> pextreal"
-  assumes "\<And>i. i \<in> P \<Longrightarrow> \<forall>n. f i n \<le> f i (Suc n)"
-  shows "(SUP n. \<Sum>i\<in>P. f i n) = (\<Sum>i\<in>P. SUP n. f i n)"
-proof cases
-  assume "finite P" from this assms show ?thesis
-  proof induct
-    case (insert i P)
-    thus ?case
-      apply simp
-      apply (subst SUPR_pextreal_add)
-      by (auto intro!: setsum_mono)
-  qed simp
-qed simp
-
-lemma psuminf_SUP_eq:
-  assumes "\<And>n i. f n i \<le> f (Suc n) i"
-  shows "(\<Sum>\<^isub>\<infinity> i. SUP n::nat. f n i) = (SUP n::nat. \<Sum>\<^isub>\<infinity> i. f n i)"
-proof -
-  { fix n :: nat
-    have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
-      using assms by (auto intro!: SUPR_pextreal_setsum[symmetric]) }
-  note * = this
-  show ?thesis
-    unfolding psuminf_def
-    unfolding *
-    apply (subst SUP_commute) ..
-qed
-
-lemma psuminf_commute:
-  shows "(\<Sum>\<^isub>\<infinity> i j. f i j) = (\<Sum>\<^isub>\<infinity> j i. f i j)"
-proof -
-  have "(SUP n. \<Sum> i < n. SUP m. \<Sum> j < m. f i j) = (SUP n. SUP m. \<Sum> i < n. \<Sum> j < m. f i j)"
-    apply (subst SUPR_pextreal_setsum)
-    by auto
-  also have "\<dots> = (SUP m n. \<Sum> j < m. \<Sum> i < n. f i j)"
-    apply (subst SUP_commute)
-    apply (subst setsum_commute)
-    by auto
-  also have "\<dots> = (SUP m. \<Sum> j < m. SUP n. \<Sum> i < n. f i j)"
-    apply (subst SUPR_pextreal_setsum)
-    by auto
-  finally show ?thesis
-    unfolding psuminf_def by auto
-qed
-
-lemma psuminf_2dimen:
-  fixes f:: "nat * nat \<Rightarrow> pextreal"
-  assumes fsums: "\<And>m. g m = (\<Sum>\<^isub>\<infinity> n. f (m,n))"
-  shows "psuminf (f \<circ> prod_decode) = psuminf g"
-proof (rule psuminf_equality)
-  fix n :: nat
-  let ?P = "prod_decode ` {..<n}"
-  have "setsum (f \<circ> prod_decode) {..<n} = setsum f ?P"
-    by (auto simp: setsum_reindex inj_prod_decode)
-  also have "\<dots> \<le> setsum f ({..Max (fst ` ?P)} \<times> {..Max (snd ` ?P)})"
-  proof (safe intro!: setsum_mono3 Max_ge image_eqI)
-    fix a b x assume "(a, b) = prod_decode x"
-    from this[symmetric] show "a = fst (prod_decode x)" "b = snd (prod_decode x)"
-      by simp_all
-  qed simp_all
-  also have "\<dots> = (\<Sum>m\<le>Max (fst ` ?P). (\<Sum>n\<le>Max (snd ` ?P). f (m,n)))"
-    unfolding setsum_cartesian_product by simp
-  also have "\<dots> \<le> (\<Sum>m\<le>Max (fst ` ?P). g m)"
-    by (auto intro!: setsum_mono psuminf_upper simp del: setsum_lessThan_Suc
-        simp: fsums lessThan_Suc_atMost[symmetric])
-  also have "\<dots> \<le> psuminf g"
-    by (auto intro!: psuminf_upper simp del: setsum_lessThan_Suc
-        simp: lessThan_Suc_atMost[symmetric])
-  finally show "setsum (f \<circ> prod_decode) {..<n} \<le> psuminf g" .
-next
-  fix y assume *: "\<And>n. setsum (f \<circ> prod_decode) {..<n} \<le> y"
-  have g: "g = (\<lambda>m. \<Sum>\<^isub>\<infinity> n. f (m,n))" unfolding fsums[symmetric] ..
-  show "psuminf g \<le> y" unfolding g
-  proof (rule psuminf_bound, unfold setsum_pinfsum[symmetric], safe intro!: psuminf_bound)
-    fix N M :: nat
-    let ?P = "{..<N} \<times> {..<M}"
-    let ?M = "Max (prod_encode ` ?P)"
-    have "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> (\<Sum>(m, n)\<in>?P. f (m, n))"
-      unfolding setsum_commute[of _ _ "{..<M}"] unfolding setsum_cartesian_product ..
-    also have "\<dots> \<le> (\<Sum>(m,n)\<in>(prod_decode ` {..?M}). f (m, n))"
-      by (auto intro!: setsum_mono3 image_eqI[where f=prod_decode, OF prod_encode_inverse[symmetric]])
-    also have "\<dots> \<le> y" using *[of "Suc ?M"]
-      by (simp add: lessThan_Suc_atMost[symmetric] setsum_reindex
-               inj_prod_decode del: setsum_lessThan_Suc)
-    finally show "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> y" .
-  qed
-qed
-
-lemma Real_max:
-  assumes "x \<ge> 0" "y \<ge> 0"
-  shows "Real (max x y) = max (Real x) (Real y)"
-  using assms unfolding max_def by (auto simp add:not_le)
-
-lemma Real_real: "Real (real x) = (if x = \<omega> then 0 else x)"
-  using assms by (cases x) auto
-
-lemma inj_on_real: "inj_on real (UNIV - {\<omega>})"
-proof (rule inj_onI)
-  fix x y assume mem: "x \<in> UNIV - {\<omega>}" "y \<in> UNIV - {\<omega>}" and "real x = real y"
-  thus "x = y" by (cases x, cases y) auto
-qed
-
-lemma inj_on_Real: "inj_on Real {0..}"
-  by (auto intro!: inj_onI)
-
-lemma range_Real[simp]: "range Real = UNIV - {\<omega>}"
-proof safe
-  fix x assume "x \<notin> range Real"
-  thus "x = \<omega>" by (cases x) auto
-qed auto
-
-lemma image_Real[simp]: "Real ` {0..} = UNIV - {\<omega>}"
-proof safe
-  fix x assume "x \<notin> Real ` {0..}"
-  thus "x = \<omega>" by (cases x) auto
-qed auto
-
-lemma pextreal_SUP_cmult:
-  fixes f :: "'a \<Rightarrow> pextreal"
-  shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)"
-proof (rule pextreal_SUPI)
-  fix i assume "i \<in> R"
-  from le_SUPI[OF this]
-  show "z * f i \<le> z * (SUP i:R. f i)" by (rule pextreal_mult_cancel)
-next
-  fix y assume "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y"
-  hence *: "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y" by auto
-  show "z * (SUP i:R. f i) \<le> y"
-  proof (cases "\<forall>i\<in>R. f i = 0")
-    case True
-    show ?thesis
-    proof cases
-      assume "R \<noteq> {}" hence "f ` R = {0}" using True by auto
-      thus ?thesis by (simp add: SUPR_def)
-    qed (simp add: SUPR_def Sup_empty bot_pextreal_def)
-  next
-    case False then obtain i where i: "i \<in> R" and f0: "f i \<noteq> 0" by auto
-    show ?thesis
-    proof (cases "z = 0 \<or> z = \<omega>")
-      case True with f0 *[OF i] show ?thesis by auto
-    next
-      case False hence z: "z \<noteq> 0" "z \<noteq> \<omega>" by auto
-      note div = pextreal_inverse_le_eq[OF this, symmetric]
-      hence "\<And>i. i\<in>R \<Longrightarrow> f i \<le> y / z" using * by auto
-      thus ?thesis unfolding div SUP_le_iff by simp
-    qed
-  qed
-qed
-
-instantiation pextreal :: topological_space
-begin
-
-definition "open A \<longleftrightarrow>
-  (\<exists>T. open T \<and> (Real ` (T\<inter>{0..}) = A - {\<omega>})) \<and> (\<omega> \<in> A \<longrightarrow> (\<exists>x\<ge>0. {Real x <..} \<subseteq> A))"
-
-lemma open_omega: "open A \<Longrightarrow> \<omega> \<in> A \<Longrightarrow> (\<exists>x\<ge>0. {Real x<..} \<subseteq> A)"
-  unfolding open_pextreal_def by auto
-
-lemma open_omegaD: assumes "open A" "\<omega> \<in> A" obtains x where "x\<ge>0" "{Real x<..} \<subseteq> A"
-  using open_omega[OF assms] by auto
-
-lemma pextreal_openE: assumes "open A" obtains A' x where
-  "open A'" "Real ` (A' \<inter> {0..}) = A - {\<omega>}"
-  "x \<ge> 0" "\<omega> \<in> A \<Longrightarrow> {Real x<..} \<subseteq> A"
-  using assms open_pextreal_def by auto
-
-instance
-proof
-  let ?U = "UNIV::pextreal set"
-  show "open ?U" unfolding open_pextreal_def
-    by (auto intro!: exI[of _ "UNIV"] exI[of _ 0])
-next
-  fix S T::"pextreal set" assume "open S" and "open T"
-  from `open S`[THEN pextreal_openE] guess S' xS . note S' = this
-  from `open T`[THEN pextreal_openE] guess T' xT . note T' = this
-
-  from S'(1-3) T'(1-3)
-  show "open (S \<inter> T)" unfolding open_pextreal_def
-  proof (safe intro!: exI[of _ "S' \<inter> T'"] exI[of _ "max xS xT"])
-    fix x assume *: "Real (max xS xT) < x" and "\<omega> \<in> S" "\<omega> \<in> T"
-    from `\<omega> \<in> S`[THEN S'(4)] * show "x \<in> S"
-      by (cases x, auto simp: max_def split: split_if_asm)
-    from `\<omega> \<in> T`[THEN T'(4)] * show "x \<in> T"
-      by (cases x, auto simp: max_def split: split_if_asm)
-  next
-    fix x assume x: "x \<notin> Real ` (S' \<inter> T' \<inter> {0..})"
-    have *: "S' \<inter> T' \<inter> {0..} = (S' \<inter> {0..}) \<inter> (T' \<inter> {0..})" by auto
-    assume "x \<in> T" "x \<in> S"
-    with S'(2) T'(2) show "x = \<omega>"
-      using x[unfolded *] inj_on_image_Int[OF inj_on_Real] by auto
-  qed auto
-next
-  fix K assume openK: "\<forall>S \<in> K. open (S:: pextreal set)"
-  hence "\<forall>S\<in>K. \<exists>T. open T \<and> Real ` (T \<inter> {0..}) = S - {\<omega>}" by (auto simp: open_pextreal_def)
-  from bchoice[OF this] guess T .. note T = this[rule_format]
-
-  show "open (\<Union>K)" unfolding open_pextreal_def
-  proof (safe intro!: exI[of _ "\<Union>(T ` K)"])
-    fix x S assume "0 \<le> x" "x \<in> T S" "S \<in> K"
-    with T[OF `S \<in> K`] show "Real x \<in> \<Union>K" by auto
-  next
-    fix x S assume x: "x \<notin> Real ` (\<Union>T ` K \<inter> {0..})" "S \<in> K" "x \<in> S"
-    hence "x \<notin> Real ` (T S \<inter> {0..})"
-      by (auto simp: image_UN UN_simps[symmetric] simp del: UN_simps)
-    thus "x = \<omega>" using T[OF `S \<in> K`] `x \<in> S` by auto
-  next
-    fix S assume "\<omega> \<in> S" "S \<in> K"
-    from openK[rule_format, OF `S \<in> K`, THEN pextreal_openE] guess S' x .
-    from this(3, 4) `\<omega> \<in> S`
-    show "\<exists>x\<ge>0. {Real x<..} \<subseteq> \<Union>K"
-      by (auto intro!: exI[of _ x] bexI[OF _ `S \<in> K`])
-  next
-    from T[THEN conjunct1] show "open (\<Union>T`K)" by auto
-  qed auto
-qed
-end
-
-lemma minus_omega[simp]: "x - \<omega> = 0" by (cases x) auto
-
-lemma open_pextreal_alt: "open A \<longleftrightarrow>
-  (\<forall>x\<in>A. \<exists>e>0. {x-e <..< x+e} \<subseteq> A) \<and> (\<omega> \<in> A \<longrightarrow> (\<exists>x\<ge>0. {Real x <..} \<subseteq> A))"
-proof -
-  have *: "(\<exists>T. open T \<and> (Real ` (T\<inter>{0..}) = A - {\<omega>})) \<longleftrightarrow>
-    open (real ` (A - {\<omega>}) \<union> {..<0})"
-  proof safe
-    fix T assume "open T" and A: "Real ` (T \<inter> {0..}) = A - {\<omega>}"
-    have *: "(\<lambda>x. real (Real x)) ` (T \<inter> {0..}) = T \<inter> {0..}"
-      by auto
-    have **: "T \<inter> {0..} \<union> {..<0} = T \<union> {..<0}" by auto
-    show "open (real ` (A - {\<omega>}) \<union> {..<0})"
-      unfolding A[symmetric] image_image * ** using `open T` by auto
-  next
-    assume "open (real ` (A - {\<omega>}) \<union> {..<0})"
-    moreover have "Real ` ((real ` (A - {\<omega>}) \<union> {..<0}) \<inter> {0..}) = A - {\<omega>}"
-      apply auto
-      apply (case_tac xb)
-      apply auto
-      apply (case_tac x)
-      apply (auto simp: image_iff)
-      apply (erule_tac x="Real r" in ballE)
-      apply auto
-      done
-    ultimately show "\<exists>T. open T \<and> Real ` (T \<inter> {0..}) = A - {\<omega>}" by auto
-  qed
-  also have "\<dots> \<longleftrightarrow> (\<forall>x\<in>A. \<exists>e>0. {x-e <..< x+e} \<subseteq> A)"
-  proof (intro iffI ballI open_subopen[THEN iffD2])
-    fix x assume *: "\<forall>x\<in>A. \<exists>e>0. {x - e<..<x + e} \<subseteq> A" and x: "x \<in> real ` (A - {\<omega>}) \<union> {..<0}"
-    show "\<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> real ` (A - {\<omega>}) \<union> {..<0}"
-    proof (cases rule: linorder_cases)
-      assume "x < 0" then show ?thesis by (intro exI[of _ "{..<0}"]) auto
-    next
-      assume "x = 0" with x
-      have "0 \<in> A"
-        apply auto by (case_tac x) auto
-      with * obtain e where "e > 0" "{0 - e<..<0 + e} \<subseteq> A" by auto
-      then have "{..<e} \<subseteq> A" using `0 \<in> A`
-        apply auto
-        apply (case_tac "x = 0")
-        by (auto dest!: pextreal_zero_lessI)
-      then have *: "{..<e} \<subseteq> A - {\<omega>}" by auto
-      def e' \<equiv> "if e = \<omega> then 1 else real e"
-      then have "0 < e'" using `e > 0` by (cases e) auto
-      have "{0..<e'} \<subseteq> real ` (A - {\<omega>})"
-      proof (cases e)
-        case infinite
-        then have "{..<e} = UNIV - {\<omega>}" by auto
-        then have A: "A - {\<omega>} = UNIV - {\<omega>}" using * by auto
-        show ?thesis unfolding e'_def infinite A
-          apply safe
-          apply (rule_tac x="Real x" in image_eqI)
-          apply auto
-          done
-      next
-        case (preal r)
-        then show ?thesis unfolding e'_def using *
-          apply safe
-          apply (rule_tac x="Real x" in image_eqI)
-          by (auto simp: subset_eq)
-      qed
-      then have "{0..<e'} \<union> {..<0} \<subseteq> real ` (A - {\<omega>}) \<union> {..<0}" by auto
-      moreover have "{0..<e'} \<union> {..<0} = {..<e'}" using `0 < e'` by auto
-      ultimately have "{..<e'} \<subseteq> real ` (A - {\<omega>}) \<union> {..<0}" by simp
-      then show ?thesis using `0 < e'` `x = 0` by auto
-    next
-      assume "0 < x"
-      with x have "Real x \<in> A" apply auto by (case_tac x) auto
-      with * obtain e where "0 < e" and e: "{Real x - e<..<Real x + e} \<subseteq> A" by auto
-      show ?thesis
-      proof cases
-        assume "e < Real x"
-        with `0 < e` obtain r where r: "e = Real r" "0 < r" by (cases e) auto
-        then have "r < x" using `e < Real x` `0 < e` by (auto split: split_if_asm)
-        then have "{x - r <..< x + r} \<subseteq> real ` (A - {\<omega>}) \<union> {..<0}"
-          using e unfolding r
-          apply (auto simp: subset_eq)
-          apply (rule_tac x="Real xa" in image_eqI)
-          by auto
-        then show ?thesis using `0 < r` by (intro exI[of _ "{x - r <..< x + r}"]) auto
-      next
-        assume "\<not> e < Real x"
-        moreover then have "Real x - e = 0" by (cases e) auto
-        moreover have "\<And>z. 0 < z \<Longrightarrow>  z * 2 < 3 * x \<Longrightarrow> Real z < Real x + e"
-          using `\<not> e < Real x` by (cases e) auto
-        ultimately have "{0 <..< x + x / 2} \<subseteq> real ` (A - {\<omega>}) \<union> {..<0}"
-          using e
-          apply (auto simp: subset_eq)
-          apply (erule_tac x="Real xa" in ballE)
-          apply (auto simp: not_less)
-          apply (rule_tac x="Real xa" in image_eqI)
-          apply auto
-          done
-        moreover have "x \<in> {0 <..< x + x / 2}" using `0 < x` by auto
-        ultimately show ?thesis by (intro exI[of _ "{0 <..< x + x / 2}"]) auto
-      qed
-    qed
-  next
-    fix x assume x: "x \<in> A" "open (real ` (A - {\<omega>}) \<union> {..<0})"
-    then show "\<exists>e>0. {x - e<..<x + e} \<subseteq> A"
-    proof (cases x)
-      case infinite then show ?thesis by (intro exI[of _ 2]) auto
-    next
-      case (preal r)
-      with `x \<in> A` have r: "r \<in> real ` (A - {\<omega>}) \<union> {..<0}" by force
-      from x(2)[unfolded open_real, THEN bspec, OF r]
-      obtain e where e: "0 < e" "\<And>x'. \<bar>x' - r\<bar> < e \<Longrightarrow> x' \<in> real ` (A - {\<omega>}) \<union> {..<0}"
-        by auto
-      show ?thesis using `0 < e` preal
-      proof (auto intro!: exI[of _ "Real e"] simp: greaterThanLessThan_iff not_less 
-                  split: split_if_asm)
-        fix z assume *: "Real (r - e) < z" "z < Real (r + e)"
-        then obtain q where [simp]: "z = Real q" "0 \<le> q" by (cases z) auto
-        with * have "r - e < q" "q < r + e" by (auto split: split_if_asm)
-        with e(2)[of q] have "q \<in> real ` (A - {\<omega>}) \<union> {..<0}" by auto
-        then show "z \<in> A" using `0 \<le> q` apply auto apply (case_tac x) apply auto done
-      next
-        fix z assume *: "0 < z" "z < Real (r + e)" "r \<le> e"
-        then obtain q where [simp]: "z = Real q" and "0 < q" by (cases z) auto
-        with * have "q < r + e" by (auto split: split_if_asm)
-        moreover have "r - e < q" using `r \<le> e` `0 < q` by auto
-        ultimately have "q \<in> real ` (A - {\<omega>}) \<union> {..<0}" using e(2)[of q] by auto
-        then show "z \<in> A" using `0 < q` apply auto apply (case_tac x) apply auto done
-      qed
-    qed
-  qed
-  finally show ?thesis unfolding open_pextreal_def by simp
-qed
-
-lemma open_pextreal_lessThan[simp]:
-  "open {..< a :: pextreal}"
-proof (cases a)
-  case (preal x) thus ?thesis unfolding open_pextreal_def
-  proof (safe intro!: exI[of _ "{..< x}"])
-    fix y assume "y < Real x"
-    moreover assume "y \<notin> Real ` ({..<x} \<inter> {0..})"
-    ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
-    thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
-  qed auto
-next
-  case infinite thus ?thesis
-    unfolding open_pextreal_def by (auto intro!: exI[of _ UNIV])
-qed
-
-lemma open_pextreal_greaterThan[simp]:
-  "open {a :: pextreal <..}"
-proof (cases a)
-  case (preal x) thus ?thesis unfolding open_pextreal_def
-  proof (safe intro!: exI[of _ "{x <..}"])
-    fix y assume "Real x < y"
-    moreover assume "y \<notin> Real ` ({x<..} \<inter> {0..})"
-    ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
-    thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
-  qed auto
-next
-  case infinite thus ?thesis
-    unfolding open_pextreal_def by (auto intro!: exI[of _ "{}"])
-qed
-
-lemma pextreal_open_greaterThanLessThan[simp]: "open {a::pextreal <..< b}"
-  unfolding greaterThanLessThan_def by auto
-
-lemma closed_pextreal_atLeast[simp, intro]: "closed {a :: pextreal ..}"
-proof -
-  have "- {a ..} = {..< a}" by auto
-  then show "closed {a ..}"
-    unfolding closed_def using open_pextreal_lessThan by auto
-qed
-
-lemma closed_pextreal_atMost[simp, intro]: "closed {.. b :: pextreal}"
-proof -
-  have "- {.. b} = {b <..}" by auto
-  then show "closed {.. b}" 
-    unfolding closed_def using open_pextreal_greaterThan by auto
-qed
-
-lemma closed_pextreal_atLeastAtMost[simp, intro]:
-  shows "closed {a :: pextreal .. b}"
-  unfolding atLeastAtMost_def by auto
-
-lemma pextreal_dense:
-  fixes x y :: pextreal assumes "x < y"
-  shows "\<exists>z. x < z \<and> z < y"
-proof -
-  from `x < y` obtain p where p: "x = Real p" "0 \<le> p" by (cases x) auto
-  show ?thesis
-  proof (cases y)
-    case (preal r) with p `x < y` have "p < r" by auto
-    with dense obtain z where "p < z" "z < r" by auto
-    thus ?thesis using preal p by (auto intro!: exI[of _ "Real z"])
-  next
-    case infinite thus ?thesis using `x < y` p
-      by (auto intro!: exI[of _ "Real p + 1"])
-  qed
-qed
-
-instance pextreal :: t2_space
-proof
-  fix x y :: pextreal assume "x \<noteq> y"
-  let "?P x (y::pextreal)" = "\<exists> U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-
-  { fix x y :: pextreal assume "x < y"
-    from pextreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
-    have "?P x y"
-      apply (rule exI[of _ "{..<z}"])
-      apply (rule exI[of _ "{z<..}"])
-      using z by auto }
-  note * = this
-
-  from `x \<noteq> y`
-  show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-  proof (cases rule: linorder_cases)
-    assume "x = y" with `x \<noteq> y` show ?thesis by simp
-  next assume "x < y" from *[OF this] show ?thesis by auto
-  next assume "y < x" from *[OF this] show ?thesis by auto
-  qed
-qed
-
-definition (in complete_lattice) isoton :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<up>" 50) where
-  "A \<up> X \<longleftrightarrow> (\<forall>i. A i \<le> A (Suc i)) \<and> (SUP i. A i) = X"
-
-definition (in complete_lattice) antiton (infix "\<down>" 50) where
-  "A \<down> X \<longleftrightarrow> (\<forall>i. A i \<ge> A (Suc i)) \<and> (INF i. A i) = X"
-
-lemma isotoneI[intro?]: "\<lbrakk> \<And>i. f i \<le> f (Suc i) ; (SUP i. f i) = F \<rbrakk> \<Longrightarrow> f \<up> F"
-  unfolding isoton_def by auto
-
-lemma (in complete_lattice) isotonD[dest]:
-  assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
-  using assms unfolding isoton_def by auto
-
-lemma isotonD'[dest]:
-  assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
-  using assms unfolding isoton_def le_fun_def by auto
-
-lemma isoton_mono_le:
-  assumes "f \<up> x" "i \<le> j"
-  shows "f i \<le> f j"
-  using `f \<up> x`[THEN isotonD(1)] lift_Suc_mono_le[of f, OF _ `i \<le> j`] by auto
-
-lemma isoton_const:
-  shows "(\<lambda> i. c) \<up> c"
-unfolding isoton_def by auto
-
-lemma isoton_cmult_right:
-  assumes "f \<up> (x::pextreal)"
-  shows "(\<lambda>i. c * f i) \<up> (c * x)"
-  using assms unfolding isoton_def pextreal_SUP_cmult
-  by (auto intro: pextreal_mult_cancel)
-
-lemma isoton_cmult_left:
-  "f \<up> (x::pextreal) \<Longrightarrow> (\<lambda>i. f i * c) \<up> (x * c)"
-  by (subst (1 2) mult_commute) (rule isoton_cmult_right)
-
-lemma isoton_add:
-  assumes "f \<up> (x::pextreal)" and "g \<up> y"
-  shows "(\<lambda>i. f i + g i) \<up> (x + y)"
-  using assms unfolding isoton_def
-  by (auto intro: pextreal_mult_cancel add_mono simp: SUPR_pextreal_add)
-
-lemma isoton_fun_expand:
-  "f \<up> x \<longleftrightarrow> (\<forall>i. (\<lambda>j. f j i) \<up> (x i))"
-proof -
-  have "\<And>i. {y. \<exists>f'\<in>range f. y = f' i} = range (\<lambda>j. f j i)"
-    by auto
-  with assms show ?thesis
-    by (auto simp add: isoton_def le_fun_def Sup_fun_def SUPR_def)
-qed
-
-lemma isoton_indicator:
-  assumes "f \<up> g"
-  shows "(\<lambda>i x. f i x * indicator A x) \<up> (\<lambda>x. g x * indicator A x :: pextreal)"
-  using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
-
-lemma isoton_setsum:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> pextreal"
-  assumes "finite A" "A \<noteq> {}"
-  assumes "\<And> x. x \<in> A \<Longrightarrow> f x \<up> y x"
-  shows "(\<lambda> i. (\<Sum> x \<in> A. f x i)) \<up> (\<Sum> x \<in> A. y x)"
-using assms
-proof (induct A rule:finite_ne_induct)
-  case singleton thus ?case by auto
-next
-  case (insert a A) note asms = this
-  hence *: "(\<lambda> i. \<Sum> x \<in> A. f x i) \<up> (\<Sum> x \<in> A. y x)" by auto
-  have **: "(\<lambda> i. f a i) \<up> y a" using asms by simp
-  have "(\<lambda> i. f a i + (\<Sum> x \<in> A. f x i)) \<up> (y a + (\<Sum> x \<in> A. y x))"
-    using * ** isoton_add by auto
-  thus "(\<lambda> i. \<Sum> x \<in> insert a A. f x i) \<up> (\<Sum> x \<in> insert a A. y x)"
-    using asms by fastsimp
-qed
-
-lemma isoton_Sup:
-  assumes "f \<up> u"
-  shows "f i \<le> u"
-  using le_SUPI[of i UNIV f] assms
-  unfolding isoton_def by auto
-
-lemma isoton_mono:
-  assumes iso: "x \<up> a" "y \<up> b" and *: "\<And>n. x n \<le> y (N n)"
-  shows "a \<le> b"
-proof -
-  from iso have "a = (SUP n. x n)" "b = (SUP n. y n)"
-    unfolding isoton_def by auto
-  with * show ?thesis by (auto intro!: SUP_mono)
-qed
-
-lemma pextreal_le_mult_one_interval:
-  fixes x y :: pextreal
-  assumes "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
-  shows "x \<le> y"
-proof (cases x, cases y)
-  assume "x = \<omega>"
-  with assms[of "1 / 2"]
-  show "x \<le> y" by simp
-next
-  fix r p assume *: "y = Real p" "x = Real r" and **: "0 \<le> r" "0 \<le> p"
-  have "r \<le> p"
-  proof (rule field_le_mult_one_interval)
-    fix z :: real assume "0 < z" and "z < 1"
-    with assms[of "Real z"]
-    show "z * r \<le> p" using ** * by (auto simp: zero_le_mult_iff)
-  qed
-  thus "x \<le> y" using ** * by simp
-qed simp
-
-lemma pextreal_greater_0[intro]:
-  fixes a :: pextreal
-  assumes "a \<noteq> 0"
-  shows "a > 0"
-using assms apply (cases a) by auto
-
-lemma pextreal_mult_strict_right_mono:
-  assumes "a < b" and "0 < c" "c < \<omega>"
-  shows "a * c < b * c"
-  using assms
-  by (cases a, cases b, cases c)
-     (auto simp: zero_le_mult_iff pextreal_less_\<omega>)
-
-lemma minus_pextreal_eq2:
-  fixes x y z :: pextreal
-  assumes "y \<le> x" and "y \<noteq> \<omega>" shows "z = x - y \<longleftrightarrow> z + y = x"
-  using assms
-  apply (subst eq_commute)
-  apply (subst minus_pextreal_eq)
-  by (cases x, cases z, auto simp add: ac_simps not_less)
-
-lemma pextreal_diff_eq_diff_imp_eq:
-  assumes "a \<noteq> \<omega>" "b \<le> a" "c \<le> a"
-  assumes "a - b = a - c"
-  shows "b = c"
-  using assms
-  by (cases a, cases b, cases c) (auto split: split_if_asm)
-
-lemma pextreal_inverse_eq_0: "inverse x = 0 \<longleftrightarrow> x = \<omega>"
-  by (cases x) auto
-
-lemma pextreal_mult_inverse:
-  "\<lbrakk> x \<noteq> \<omega> ; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x * inverse x = 1"
-  by (cases x) auto
-
-lemma pextreal_zero_less_diff_iff:
-  fixes a b :: pextreal shows "0 < a - b \<longleftrightarrow> b < a"
-  apply (cases a, cases b)
-  apply (auto simp: pextreal_noteq_omega_Ex pextreal_less_\<omega>)
-  apply (cases b)
-  by auto
-
-lemma pextreal_less_Real_Ex:
-  fixes a b :: pextreal shows "x < Real r \<longleftrightarrow> (\<exists>p\<ge>0. p < r \<and> x = Real p)"
-  by (cases x) auto
-
-lemma open_Real: assumes "open S" shows "open (Real ` ({0..} \<inter> S))"
-  unfolding open_pextreal_def apply(rule,rule,rule,rule assms) by auto
-
-lemma pextreal_zero_le_diff:
-  fixes a b :: pextreal shows "a - b = 0 \<longleftrightarrow> a \<le> b"
-  by (cases a, cases b, simp_all, cases b, auto)
-
-lemma lim_Real[simp]: assumes "\<forall>n. f n \<ge> 0" "m\<ge>0"
-  shows "(\<lambda>n. Real (f n)) ----> Real m \<longleftrightarrow> (\<lambda>n. f n) ----> m" (is "?l = ?r")
-proof assume ?l show ?r unfolding Lim_sequentially
-  proof safe fix e::real assume e:"e>0"
-    note open_ball[of m e] note open_Real[OF this]
-    note * = `?l`[unfolded tendsto_def,rule_format,OF this]
-    have "eventually (\<lambda>x. Real (f x) \<in> Real ` ({0..} \<inter> ball m e)) sequentially"
-      apply(rule *) unfolding image_iff using assms(2) e by auto
-    thus "\<exists>N. \<forall>n\<ge>N. dist (f n) m < e" unfolding eventually_sequentially 
-      apply safe apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
-    proof- fix n x assume "Real (f n) = Real x" "0 \<le> x"
-      hence *:"f n = x" using assms(1) by auto
-      assume "x \<in> ball m e" thus "dist (f n) m < e" unfolding *
-        by (auto simp add:dist_commute)
-    qed qed
-next assume ?r show ?l unfolding tendsto_def eventually_sequentially 
-  proof safe fix S assume S:"open S" "Real m \<in> S"
-    guess T y using S(1) apply-apply(erule pextreal_openE) . note T=this
-    have "m\<in>real ` (S - {\<omega>})" unfolding image_iff 
-      apply(rule_tac x="Real m" in bexI) using assms(2) S(2) by auto
-    hence "m \<in> T" unfolding T(2)[THEN sym] by auto 
-    from `?r`[unfolded tendsto_def eventually_sequentially,rule_format,OF T(1) this]
-    guess N .. note N=this[rule_format]
-    show "\<exists>N. \<forall>n\<ge>N. Real (f n) \<in> S" apply(rule_tac x=N in exI) 
-    proof safe fix n assume n:"N\<le>n"
-      have "f n \<in> real ` (S - {\<omega>})" using N[OF n] assms unfolding T(2)[THEN sym] 
-        unfolding image_iff apply-apply(rule_tac x="Real (f n)" in bexI)
-        unfolding real_Real by auto
-      then guess x unfolding image_iff .. note x=this
-      show "Real (f n) \<in> S" unfolding x apply(subst Real_real) using x by auto
-    qed
-  qed
-qed
-
-lemma pextreal_INFI:
-  fixes x :: pextreal
-  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
-  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x"
-  shows "(INF i:A. f i) = x"
-  unfolding INFI_def Inf_pextreal_def
-  using assms by (auto intro!: Greatest_equality)
-
-lemma real_of_pextreal_less:"x < y \<Longrightarrow> y\<noteq>\<omega> \<Longrightarrow> real x < real y"
-proof- case goal1
-  have *:"y = Real (real y)" "x = Real (real x)" using goal1 Real_real by auto
-  show ?case using goal1 apply- apply(subst(asm) *(1))apply(subst(asm) *(2))
-    unfolding pextreal_less by auto
-qed
-
-lemma not_less_omega[simp]:"\<not> x < \<omega> \<longleftrightarrow> x = \<omega>"
-  by (metis antisym_conv3 pextreal_less(3)) 
-
-lemma Real_real': assumes "x\<noteq>\<omega>" shows "Real (real x) = x"
-proof- have *:"(THE r. 0 \<le> r \<and> x = Real r) = real x"
-    apply(rule the_equality) using assms unfolding Real_real by auto
-  have "Real (THE r. 0 \<le> r \<and> x = Real r) = x" unfolding *
-    using assms unfolding Real_real by auto
-  thus ?thesis unfolding real_of_pextreal_def of_pextreal_def
-    unfolding pextreal_case_def using assms by auto
-qed 
-
-lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)" 
-  unfolding pextreal_less by auto
-
-lemma Lim_omega: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
-proof assume ?r show ?l apply(rule topological_tendstoI)
-    unfolding eventually_sequentially
-  proof- fix S assume "open S" "\<omega> \<in> S"
-    from open_omega[OF this] guess B .. note B=this
-    from `?r`[rule_format,of "(max B 0)+1"] guess N .. note N=this
-    show "\<exists>N. \<forall>n\<ge>N. f n \<in> S" apply(rule_tac x=N in exI)
-    proof safe case goal1 
-      have "Real B < Real ((max B 0) + 1)" by auto
-      also have "... \<le> f n" using goal1 N by auto
-      finally show ?case using B by fastsimp
-    qed
-  qed
-next assume ?l show ?r
-  proof fix B::real have "open {Real B<..}" "\<omega> \<in> {Real B<..}" by auto
-    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
-    guess N .. note N=this
-    show "\<exists>N. \<forall>n\<ge>N. Real B \<le> f n" apply(rule_tac x=N in exI) using N by auto
-  qed
-qed
-
-lemma Lim_bounded_omgea: assumes lim:"f ----> l" and "\<And>n. f n \<le> Real B" shows "l \<noteq> \<omega>"
-proof(rule ccontr,unfold not_not) let ?B = "max (B + 1) 1" assume as:"l=\<omega>"
-  from lim[unfolded this Lim_omega,rule_format,of "?B"]
-  guess N .. note N=this[rule_format,OF le_refl]
-  hence "Real ?B \<le> Real B" using assms(2)[of N] by(rule order_trans) 
-  hence "Real ?B < Real ?B" using Real_less_plus_one[of B] by(rule le_less_trans)
-  thus False by auto
-qed
-
-lemma incseq_le_pextreal: assumes inc: "\<And>n m. n\<ge>m \<Longrightarrow> X n \<ge> X m"
-  and lim: "X ----> (L::pextreal)" shows "X n \<le> L"
-proof(cases "L = \<omega>")
-  case False have "\<forall>n. X n \<noteq> \<omega>"
-  proof(rule ccontr,unfold not_all not_not,safe)
-    case goal1 hence "\<forall>n\<ge>x. X n = \<omega>" using inc[of x] by auto
-    hence "X ----> \<omega>" unfolding tendsto_def eventually_sequentially
-      apply safe apply(rule_tac x=x in exI) by auto
-    note Lim_unique[OF trivial_limit_sequentially this lim]
-    with False show False by auto
-  qed note * =this[rule_format]
-
-  have **:"\<forall>m n. m \<le> n \<longrightarrow> Real (real (X m)) \<le> Real (real (X n))"
-    unfolding Real_real using * inc by auto
-  have "real (X n) \<le> real L" apply-apply(rule incseq_le) defer
-    apply(subst lim_Real[THEN sym]) apply(rule,rule,rule)
-    unfolding Real_real'[OF *] Real_real'[OF False] 
-    unfolding incseq_def using ** lim by auto
-  hence "Real (real (X n)) \<le> Real (real L)" by auto
-  thus ?thesis unfolding Real_real using * False by auto
-qed auto
-
-lemma SUP_Lim_pextreal: assumes "\<And>n m. n\<ge>m \<Longrightarrow> f n \<ge> f m" "f ----> l"
-  shows "(SUP n. f n) = (l::pextreal)" unfolding SUPR_def Sup_pextreal_def
-proof (safe intro!: Least_equality)
-  fix n::nat show "f n \<le> l" apply(rule incseq_le_pextreal)
-    using assms by auto
-next fix y assume y:"\<forall>x\<in>range f. x \<le> y" show "l \<le> y"
-  proof(rule ccontr,cases "y=\<omega>",unfold not_le)
-    case False assume as:"y < l"
-    have l:"l \<noteq> \<omega>" apply(rule Lim_bounded_omgea[OF assms(2), of "real y"])
-      using False y unfolding Real_real by auto
-
-    have yl:"real y < real l" using as apply-
-      apply(subst(asm) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
-      apply(subst(asm) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) 
-      unfolding pextreal_less apply(subst(asm) if_P) by auto
-    hence "y + (y - l) * Real (1 / 2) < l" apply-
-      apply(subst Real_real'[THEN sym,OF `y\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
-      apply(subst Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) by auto
-    hence *:"l \<in> {y + (y - l) / 2<..}" by auto
-    have "open {y + (y-l)/2 <..}" by auto
-    note topological_tendstoD[OF assms(2) this *]
-    from this[unfolded eventually_sequentially] guess N .. note this[rule_format, of N]
-    hence "y + (y - l) * Real (1 / 2) < y" using y[rule_format,of "f N"] by auto
-    hence "Real (real y) + (Real (real y) - Real (real l)) * Real (1 / 2) < Real (real y)"
-      unfolding Real_real using `y\<noteq>\<omega>` `l\<noteq>\<omega>` by auto
-    thus False using yl by auto
-  qed auto
-qed
-
-lemma Real_max':"Real x = Real (max x 0)" 
-proof(cases "x < 0") case True
-  hence *:"max x 0 = 0" by auto
-  show ?thesis unfolding * using True by auto
-qed auto
-
-lemma lim_pextreal_increasing: assumes "\<forall>n m. n\<ge>m \<longrightarrow> f n \<ge> f m"
-  obtains l where "f ----> (l::pextreal)"
-proof(cases "\<exists>B. \<forall>n. f n < Real B")
-  case False thus thesis apply- apply(rule that[of \<omega>]) unfolding Lim_omega not_ex not_all
-    apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
-    apply(rule order_trans[OF _ assms[rule_format]]) by auto
-next case True then guess B .. note B = this[rule_format]
-  hence *:"\<And>n. f n < \<omega>" apply-apply(rule less_le_trans,assumption) by auto
-  have *:"\<And>n. f n \<noteq> \<omega>" proof- case goal1 show ?case using *[of n] by auto qed
-  have B':"\<And>n. real (f n) \<le> max 0 B" proof- case goal1 thus ?case
-      using B[of n] apply-apply(subst(asm) Real_real'[THEN sym]) defer
-      apply(subst(asm)(2) Real_max') unfolding pextreal_less apply(subst(asm) if_P) using *[of n] by auto
-  qed
-  have "\<exists>l. (\<lambda>n. real (f n)) ----> l" apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
-  proof safe show "bounded {real (f n) |n. True}"
-      unfolding bounded_def apply(rule_tac x=0 in exI,rule_tac x="max 0 B" in exI)
-      using B' unfolding dist_norm by auto
-    fix n::nat have "Real (real (f n)) \<le> Real (real (f (Suc n)))"
-      using assms[rule_format,of n "Suc n"] apply(subst Real_real)+
-      using *[of n] *[of "Suc n"] by fastsimp
-    thus "real (f n) \<le> real (f (Suc n))" by auto
-  qed then guess l .. note l=this
-  have "0 \<le> l" apply(rule LIMSEQ_le_const[OF l])
-    by(rule_tac x=0 in exI,auto)
-
-  thus ?thesis apply-apply(rule that[of "Real l"])
-    using l apply-apply(subst(asm) lim_Real[THEN sym]) prefer 3
-    unfolding Real_real using * by auto
-qed
-
-lemma setsum_neq_omega: assumes "finite s" "\<And>x. x \<in> s \<Longrightarrow> f x \<noteq> \<omega>"
-  shows "setsum f s \<noteq> \<omega>" using assms
-proof induct case (insert x s)
-  show ?case unfolding setsum.insert[OF insert(1-2)] 
-    using insert by auto
-qed auto
-
-
-lemma real_Real': "0 \<le> x \<Longrightarrow> real (Real x) = x"
-  unfolding real_Real by auto
-
-lemma real_pextreal_pos[intro]:
-  assumes "x \<noteq> 0" "x \<noteq> \<omega>"
-  shows "real x > 0"
-  apply(subst real_Real'[THEN sym,of 0]) defer
-  apply(rule real_of_pextreal_less) using assms by auto
-
-lemma Lim_omega_gt: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n > Real B)" (is "?l = ?r")
-proof assume ?l thus ?r unfolding Lim_omega apply safe
-    apply(erule_tac x="max B 0 +1" in allE,safe)
-    apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
-    apply(rule_tac y="Real (max B 0 + 1)" in less_le_trans) by auto
-next assume ?r thus ?l unfolding Lim_omega apply safe
-    apply(erule_tac x=B in allE,safe) apply(rule_tac x=N in exI,safe) by auto
-qed
-
-lemma pextreal_minus_le_cancel:
-  fixes a b c :: pextreal
-  assumes "b \<le> a"
-  shows "c - a \<le> c - b"
-  using assms by (cases a, cases b, cases c, simp, simp, simp, cases b, cases c, simp_all)
-
-lemma pextreal_minus_\<omega>[simp]: "x - \<omega> = 0" by (cases x) simp_all
-
-lemma pextreal_minus_mono[intro]: "a - x \<le> (a::pextreal)"
-proof- have "a - x \<le> a - 0"
-    apply(rule pextreal_minus_le_cancel) by auto
-  thus ?thesis by auto
-qed
-
-lemma pextreal_minus_eq_\<omega>[simp]: "x - y = \<omega> \<longleftrightarrow> (x = \<omega> \<and> y \<noteq> \<omega>)"
-  by (cases x, cases y) (auto, cases y, auto)
-
-lemma pextreal_less_minus_iff:
-  fixes a b c :: pextreal
-  shows "a < b - c \<longleftrightarrow> c + a < b"
-  by (cases c, cases a, cases b, auto)
-
-lemma pextreal_minus_less_iff:
-  fixes a b c :: pextreal shows "a - c < b \<longleftrightarrow> (0 < b \<and> (c \<noteq> \<omega> \<longrightarrow> a < b + c))"
-  by (cases c, cases a, cases b, auto)
-
-lemma pextreal_le_minus_iff:
-  fixes a b c :: pextreal
-  shows "a \<le> c - b \<longleftrightarrow> ((c \<le> b \<longrightarrow> a = 0) \<and> (b < c \<longrightarrow> a + b \<le> c))"
-  by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex)
-
-lemma pextreal_minus_le_iff:
-  fixes a b c :: pextreal
-  shows "a - c \<le> b \<longleftrightarrow> (c \<le> a \<longrightarrow> a \<le> b + c)"
-  by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex)
-
-lemmas pextreal_minus_order = pextreal_minus_le_iff pextreal_minus_less_iff pextreal_le_minus_iff pextreal_less_minus_iff
-
-lemma pextreal_minus_strict_mono:
-  assumes "a > 0" "x > 0" "a\<noteq>\<omega>"
-  shows "a - x < (a::pextreal)"
-  using assms by(cases x, cases a, auto)
-
-lemma pextreal_minus':
-  "Real r - Real p = (if 0 \<le> r \<and> p \<le> r then if 0 \<le> p then Real (r - p) else Real r else 0)"
-  by (auto simp: minus_pextreal_eq not_less)
-
-lemma pextreal_minus_plus:
-  "x \<le> (a::pextreal) \<Longrightarrow> a - x + x = a"
-  by (cases a, cases x) auto
-
-lemma pextreal_cancel_plus_minus: "b \<noteq> \<omega> \<Longrightarrow> a + b - b = a"
-  by (cases a, cases b) auto
-
-lemma pextreal_minus_le_cancel_right:
-  fixes a b c :: pextreal
-  assumes "a \<le> b" "c \<le> a"
-  shows "a - c \<le> b - c"
-  using assms by (cases a, cases b, cases c, auto, cases c, auto)
-
-lemma real_of_pextreal_setsum':
-  assumes "\<forall>x \<in> S. f x \<noteq> \<omega>"
-  shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
-proof cases
-  assume "finite S"
-  from this assms show ?thesis
-    by induct (simp_all add: real_of_pextreal_add setsum_\<omega>)
-qed simp
-
-lemma Lim_omega_pos: "f ----> \<omega> \<longleftrightarrow> (\<forall>B>0. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
-  unfolding Lim_omega apply safe defer
-  apply(erule_tac x="max 1 B" in allE) apply safe defer
-  apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
-  apply(rule_tac y="Real (max 1 B)" in order_trans) by auto
-
-lemma pextreal_LimI_finite:
-  assumes "x \<noteq> \<omega>" "\<And>r. 0 < r \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
-  shows "u ----> x"
-proof (rule topological_tendstoI, unfold eventually_sequentially)
-  fix S assume "open S" "x \<in> S"
-  then obtain A where "open A" and A_eq: "Real ` (A \<inter> {0..}) = S - {\<omega>}" by (auto elim!: pextreal_openE)
-  then have "x \<in> Real ` (A \<inter> {0..})" using `x \<in> S` `x \<noteq> \<omega>` by auto
-  then have "real x \<in> A" by auto
-  then obtain r where "0 < r" and dist: "\<And>y. dist y (real x) < r \<Longrightarrow> y \<in> A"
-    using `open A` unfolding open_real_def by auto
-  then obtain n where
-    upper: "\<And>N. n \<le> N \<Longrightarrow> u N < x + Real r" and
-    lower: "\<And>N. n \<le> N \<Longrightarrow> x < u N + Real r" using assms(2)[of "Real r"] by auto
-  show "\<exists>N. \<forall>n\<ge>N. u n \<in> S"
-  proof (safe intro!: exI[of _ n])
-    fix N assume "n \<le> N"
-    from upper[OF this] `x \<noteq> \<omega>` `0 < r`
-    have "u N \<noteq> \<omega>" by (force simp: pextreal_noteq_omega_Ex)
-    with `x \<noteq> \<omega>` `0 < r` lower[OF `n \<le> N`] upper[OF `n \<le> N`]
-    have "dist (real (u N)) (real x) < r" "u N \<noteq> \<omega>"
-      by (auto simp: pextreal_noteq_omega_Ex dist_real_def abs_diff_less_iff field_simps)
-    from dist[OF this(1)]
-    have "u N \<in> Real ` (A \<inter> {0..})" using `u N \<noteq> \<omega>`
-      by (auto intro!: image_eqI[of _ _ "real (u N)"] simp: pextreal_noteq_omega_Ex Real_real)
-    thus "u N \<in> S" using A_eq by simp
-  qed
-qed
-
-lemma real_Real_max:"real (Real x) = max x 0"
-  unfolding real_Real by auto
-
-lemma Sup_lim:
-  assumes "\<forall>n. b n \<in> s" "b ----> (a::pextreal)"
-  shows "a \<le> Sup s"
-proof(rule ccontr,unfold not_le)
-  assume as:"Sup s < a" hence om:"Sup s \<noteq> \<omega>" by auto
-  have s:"s \<noteq> {}" using assms by auto
-  { presume *:"\<forall>n. b n < a \<Longrightarrow> False"
-    show False apply(cases,rule *,assumption,unfold not_all not_less)
-    proof- case goal1 then guess n .. note n=this
-      thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]]
-        using as by auto
-    qed
-  } assume b:"\<forall>n. b n < a"
-  show False
-  proof(cases "a = \<omega>")
-    case False have *:"a - Sup s > 0" 
-      using False as by(auto simp: pextreal_zero_le_diff)
-    have "(a - Sup s) / 2 \<le> a / 2" unfolding divide_pextreal_def
-      apply(rule mult_right_mono) by auto
-    also have "... = Real (real (a / 2))" apply(rule Real_real'[THEN sym])
-      using False by auto
-    also have "... < Real (real a)" unfolding pextreal_less using as False
-      by(auto simp add: real_of_pextreal_mult[THEN sym])
-    also have "... = a" apply(rule Real_real') using False by auto
-    finally have asup:"a > (a - Sup s) / 2" .
-    have "\<exists>n. a - b n < (a - Sup s) / 2"
-    proof(rule ccontr,unfold not_ex not_less)
-      case goal1
-      have "(a - Sup s) * Real (1 / 2)  > 0" 
-        using * by auto
-      hence "a - (a - Sup s) * Real (1 / 2) < a"
-        apply-apply(rule pextreal_minus_strict_mono)
-        using False * by auto
-      hence *:"a \<in> {a - (a - Sup s) / 2<..}"using asup by auto 
-      note topological_tendstoD[OF assms(2) open_pextreal_greaterThan,OF *]
-      from this[unfolded eventually_sequentially] guess n .. 
-      note n = this[rule_format,of n] 
-      have "b n + (a - Sup s) / 2 \<le> a" 
-        using add_right_mono[OF goal1[rule_format,of n],of "b n"]
-        unfolding pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]]
-        by(auto simp: add_commute)
-      hence "b n \<le> a - (a - Sup s) / 2" unfolding pextreal_le_minus_iff
-        using asup by auto
-      hence "b n \<notin> {a - (a - Sup s) / 2<..}" by auto
-      thus False using n by auto
-    qed
-    then guess n .. note n = this
-    have "Sup s < a - (a - Sup s) / 2"
-      using False as om by (cases a) (auto simp: pextreal_noteq_omega_Ex field_simps)
-    also have "... \<le> b n"
-    proof- note add_right_mono[OF less_imp_le[OF n],of "b n"]
-      note this[unfolded pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]]]
-      hence "a - (a - Sup s) / 2 \<le> (a - Sup s) / 2 + b n - (a - Sup s) / 2"
-        apply(rule pextreal_minus_le_cancel_right) using asup by auto
-      also have "... = b n + (a - Sup s) / 2 - (a - Sup s) / 2" 
-        by(auto simp add: add_commute)
-      also have "... = b n" apply(subst pextreal_cancel_plus_minus)
-      proof(rule ccontr,unfold not_not) case goal1
-        show ?case using asup unfolding goal1 by auto 
-      qed auto
-      finally show ?thesis .
-    qed     
-    finally show False
-      using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] by auto  
-  next case True
-    from assms(2)[unfolded True Lim_omega_gt,rule_format,of "real (Sup s)"]
-    guess N .. note N = this[rule_format,of N]
-    thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of N]] 
-      unfolding Real_real using om by auto
-  qed qed
-
-lemma Sup_mono_lim:
-  assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> b ----> (a::pextreal)"
-  shows "Sup A \<le> Sup B"
-  unfolding Sup_le_iff apply(rule) apply(drule assms[rule_format]) apply safe
-  apply(rule_tac b=b in Sup_lim) by auto
-
-lemma pextreal_less_add:
-  assumes "x \<noteq> \<omega>" "a < b"
-  shows "x + a < x + b"
-  using assms by (cases a, cases b, cases x) auto
-
-lemma SUPR_lim:
-  assumes "\<forall>n. b n \<in> B" "(\<lambda>n. f (b n)) ----> (f a::pextreal)"
-  shows "f a \<le> SUPR B f"
-  unfolding SUPR_def apply(rule Sup_lim[of "\<lambda>n. f (b n)"])
-  using assms by auto
-
-lemma SUP_\<omega>_imp:
-  assumes "(SUP i. f i) = \<omega>"
-  shows "\<exists>i. Real x < f i"
-proof (rule ccontr)
-  assume "\<not> ?thesis" hence "\<And>i. f i \<le> Real x" by (simp add: not_less)
-  hence "(SUP i. f i) \<le> Real x" unfolding SUP_le_iff by auto
-  with assms show False by auto
-qed
-
-lemma SUPR_mono_lim:
-  assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> (\<lambda>n. f (b n)) ----> (f a::pextreal)"
-  shows "SUPR A f \<le> SUPR B f"
-  unfolding SUPR_def apply(rule Sup_mono_lim)
-  apply safe apply(drule assms[rule_format],safe)
-  apply(rule_tac x="\<lambda>n. f (b n)" in exI) by auto
-
-lemma real_0_imp_eq_0:
-  assumes "x \<noteq> \<omega>" "real x = 0"
-  shows "x = 0"
-using assms by (cases x) auto
-
-lemma SUPR_mono:
-  assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
-  shows "SUPR A f \<le> SUPR B f"
-  unfolding SUPR_def apply(rule Sup_mono)
-  using assms by auto
-
-lemma less_add_Real:
-  fixes x :: real
-  fixes a b :: pextreal
-  assumes "x \<ge> 0" "a < b"
-  shows "a + Real x < b + Real x"
-using assms by (cases a, cases b) auto
-
-lemma le_add_Real:
-  fixes x :: real
-  fixes a b :: pextreal
-  assumes "x \<ge> 0" "a \<le> b"
-  shows "a + Real x \<le> b + Real x"
-using assms by (cases a, cases b) auto
-
-lemma le_imp_less_pextreal:
-  fixes x :: pextreal
-  assumes "x > 0" "a + x \<le> b" "a \<noteq> \<omega>"
-  shows "a < b"
-using assms by (cases x, cases a, cases b) auto
-
-lemma pextreal_INF_minus:
-  fixes f :: "nat \<Rightarrow> pextreal"
-  assumes "c \<noteq> \<omega>"
-  shows "(INF i. c - f i) = c - (SUP i. f i)"
-proof (cases "SUP i. f i")
-  case infinite
-  from `c \<noteq> \<omega>` obtain x where [simp]: "c = Real x" by (cases c) auto
-  from SUP_\<omega>_imp[OF infinite] obtain i where "Real x < f i" by auto
-  have "(INF i. c - f i) \<le> c - f i"
-    by (auto intro!: complete_lattice_class.INF_leI)
-  also have "\<dots> = 0" using `Real x < f i` by (auto simp: minus_pextreal_eq)
-  finally show ?thesis using infinite by auto
-next
-  case (preal r)
-  from `c \<noteq> \<omega>` obtain x where c: "c = Real x" by (cases c) auto
-
-  show ?thesis unfolding c
-  proof (rule pextreal_INFI)
-    fix i have "f i \<le> (SUP i. f i)" by (rule le_SUPI) simp
-    thus "Real x - (SUP i. f i) \<le> Real x - f i" by (rule pextreal_minus_le_cancel)
-  next
-    fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> y \<le> Real x - f i"
-    from this[of 0] obtain p where p: "y = Real p" "0 \<le> p"
-      by (cases "f 0", cases y, auto split: split_if_asm)
-    hence "\<And>i. Real p \<le> Real x - f i" using * by auto
-    hence *: "\<And>i. Real x \<le> f i \<Longrightarrow> Real p = 0"
-      "\<And>i. f i < Real x \<Longrightarrow> Real p + f i \<le> Real x"
-      unfolding pextreal_le_minus_iff by auto
-    show "y \<le> Real x - (SUP i. f i)" unfolding p pextreal_le_minus_iff
-    proof safe
-      assume x_less: "Real x \<le> (SUP i. f i)"
-      show "Real p = 0"
-      proof (rule ccontr)
-        assume "Real p \<noteq> 0"
-        hence "0 < Real p" by auto
-        from Sup_close[OF this, of "range f"]
-        obtain i where e: "(SUP i. f i) < f i + Real p"
-          using preal unfolding SUPR_def by auto
-        hence "Real x \<le> f i + Real p" using x_less by auto
-        show False
-        proof cases
-          assume "\<forall>i. f i < Real x"
-          hence "Real p + f i \<le> Real x" using * by auto
-          hence "f i + Real p \<le> (SUP i. f i)" using x_less by (auto simp: field_simps)
-          thus False using e by auto
-        next
-          assume "\<not> (\<forall>i. f i < Real x)"
-          then obtain i where "Real x \<le> f i" by (auto simp: not_less)
-          from *(1)[OF this] show False using `Real p \<noteq> 0` by auto
-        qed
-      qed
-    next
-      have "\<And>i. f i \<le> (SUP i. f i)" by (rule complete_lattice_class.le_SUPI) auto
-      also assume "(SUP i. f i) < Real x"
-      finally have "\<And>i. f i < Real x" by auto
-      hence *: "\<And>i. Real p + f i \<le> Real x" using * by auto
-      have "Real p \<le> Real x" using *[of 0] by (cases "f 0") (auto split: split_if_asm)
-
-      have SUP_eq: "(SUP i. f i) \<le> Real x - Real p"
-      proof (rule SUP_leI)
-        fix i show "f i \<le> Real x - Real p" unfolding pextreal_le_minus_iff
-        proof safe
-          assume "Real x \<le> Real p"
-          with *[of i] show "f i = 0"
-            by (cases "f i") (auto split: split_if_asm)
-        next
-          assume "Real p < Real x"
-          show "f i + Real p \<le> Real x" using * by (auto simp: field_simps)
-        qed
-      qed
-
-      show "Real p + (SUP i. f i) \<le> Real x"
-      proof cases
-        assume "Real x \<le> Real p"
-        with `Real p \<le> Real x` have [simp]: "Real p = Real x" by (rule antisym)
-        { fix i have "f i = 0" using *[of i] by (cases "f i") (auto split: split_if_asm) }
-        hence "(SUP i. f i) \<le> 0" by (auto intro!: SUP_leI)
-        thus ?thesis by simp
-      next
-        assume "\<not> Real x \<le> Real p" hence "Real p < Real x" unfolding not_le .
-        with SUP_eq show ?thesis unfolding pextreal_le_minus_iff by (auto simp: field_simps)
-      qed
-    qed
-  qed
-qed
-
-lemma pextreal_SUP_minus:
-  fixes f :: "nat \<Rightarrow> pextreal"
-  shows "(SUP i. c - f i) = c - (INF i. f i)"
-proof (rule pextreal_SUPI)
-  fix i have "(INF i. f i) \<le> f i" by (rule INF_leI) simp
-  thus "c - f i \<le> c - (INF i. f i)" by (rule pextreal_minus_le_cancel)
-next
-  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c - f i \<le> y"
-  show "c - (INF i. f i) \<le> y"
-  proof (cases y)
-    case (preal p)
-
-    show ?thesis unfolding pextreal_minus_le_iff preal
-    proof safe
-      assume INF_le_x: "(INF i. f i) \<le> c"
-      from * have *: "\<And>i. f i \<le> c \<Longrightarrow> c \<le> Real p + f i"
-        unfolding pextreal_minus_le_iff preal by auto
-
-      have INF_eq: "c - Real p \<le> (INF i. f i)"
-      proof (rule le_INFI)
-        fix i show "c - Real p \<le> f i" unfolding pextreal_minus_le_iff
-        proof safe
-          assume "Real p \<le> c"
-          show "c \<le> f i + Real p"
-          proof cases
-            assume "f i \<le> c" from *[OF this]
-            show ?thesis by (simp add: field_simps)
-          next
-            assume "\<not> f i \<le> c"
-            hence "c \<le> f i" by auto
-            also have "\<dots> \<le> f i + Real p" by auto
-            finally show ?thesis .
-          qed
-        qed
-      qed
-
-      show "c \<le> Real p + (INF i. f i)"
-      proof cases
-        assume "Real p \<le> c"
-        with INF_eq show ?thesis unfolding pextreal_minus_le_iff by (auto simp: field_simps)
-      next
-        assume "\<not> Real p \<le> c"
-        hence "c \<le> Real p" by auto
-        also have "Real p \<le> Real p + (INF i. f i)" by auto
-        finally show ?thesis .
-      qed
-    qed
-  qed simp
-qed
-
-lemma pextreal_le_minus_imp_0:
-  fixes a b :: pextreal
-  shows "a \<le> a - b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a \<noteq> \<omega> \<Longrightarrow> b = 0"
-  by (cases a, cases b, auto split: split_if_asm)
-
-lemma lim_INF_eq_lim_SUP:
-  fixes X :: "nat \<Rightarrow> real"
-  assumes "\<And>i. 0 \<le> X i" and "0 \<le> x"
-  and lim_INF: "(SUP n. INF m. Real (X (n + m))) = Real x" (is "(SUP n. ?INF n) = _")
-  and lim_SUP: "(INF n. SUP m. Real (X (n + m))) = Real x" (is "(INF n. ?SUP n) = _")
-  shows "X ----> x"
-proof (rule LIMSEQ_I)
-  fix r :: real assume "0 < r"
-  hence "0 \<le> r" by auto
-  from Sup_close[of "Real r" "range ?INF"]
-  obtain n where inf: "Real x < ?INF n + Real r"
-    unfolding SUPR_def lim_INF[unfolded SUPR_def] using `0 < r` by auto
-
-  from Inf_close[of "range ?SUP" "Real r"]
-  obtain n' where sup: "?SUP n' < Real x + Real r"
-    unfolding INFI_def lim_SUP[unfolded INFI_def] using `0 < r` by auto
-
-  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
-  proof (safe intro!: exI[of _ "max n n'"])
-    fix m assume "max n n' \<le> m" hence "n \<le> m" "n' \<le> m" by auto
-
-    note inf
-    also have "?INF n + Real r \<le> Real (X (n + (m - n))) + Real r"
-      by (rule le_add_Real, auto simp: `0 \<le> r` intro: INF_leI)
-    finally have up: "x < X m + r"
-      using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n \<le> m` by auto
-
-    have "Real (X (n' + (m - n'))) \<le> ?SUP n'"
-      by (auto simp: `0 \<le> r` intro: le_SUPI)
-    also note sup
-    finally have down: "X m < x + r"
-      using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n' \<le> m` by auto
-
-    show "norm (X m - x) < r" using up down by auto
-  qed
-qed
-
-lemma Sup_countable_SUPR:
-  assumes "Sup A \<noteq> \<omega>" "A \<noteq> {}"
-  shows "\<exists> f::nat \<Rightarrow> pextreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
-proof -
-  have "\<And>n. 0 < 1 / (of_nat n :: pextreal)" by auto
-  from Sup_close[OF this assms]
-  have "\<forall>n. \<exists>x. x \<in> A \<and> Sup A < x + 1 / of_nat n" by blast
-  from choice[OF this] obtain f where "range f \<subseteq> A" and
-    epsilon: "\<And>n. Sup A < f n + 1 / of_nat n" by blast
-  have "SUPR UNIV f = Sup A"
-  proof (rule pextreal_SUPI)
-    fix i show "f i \<le> Sup A" using `range f \<subseteq> A`
-      by (auto intro!: complete_lattice_class.Sup_upper)
-  next
-    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
-    show "Sup A \<le> y"
-    proof (rule pextreal_le_epsilon)
-      fix e :: pextreal assume "0 < e"
-      show "Sup A \<le> y + e"
-      proof (cases e)
-        case (preal r)
-        hence "0 < r" using `0 < e` by auto
-        then obtain n where *: "inverse (of_nat n) < r" "0 < n"
-          using ex_inverse_of_nat_less by auto
-        have "Sup A \<le> f n + 1 / of_nat n" using epsilon[of n] by auto
-        also have "1 / of_nat n \<le> e" using preal * by (auto simp: real_eq_of_nat)
-        with bound have "f n + 1 / of_nat n \<le> y + e" by (rule add_mono) simp
-        finally show "Sup A \<le> y + e" .
-      qed simp
-    qed
-  qed
-  with `range f \<subseteq> A` show ?thesis by (auto intro!: exI[of _ f])
-qed
-
-lemma SUPR_countable_SUPR:
-  assumes "SUPR A g \<noteq> \<omega>" "A \<noteq> {}"
-  shows "\<exists> f::nat \<Rightarrow> pextreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
-proof -
-  have "Sup (g`A) \<noteq> \<omega>" "g`A \<noteq> {}" using assms unfolding SUPR_def by auto
-  from Sup_countable_SUPR[OF this]
-  show ?thesis unfolding SUPR_def .
-qed
-
-lemma pextreal_setsum_subtractf:
-  assumes "\<And>i. i\<in>A \<Longrightarrow> g i \<le> f i" and "\<And>i. i\<in>A \<Longrightarrow> f i \<noteq> \<omega>"
-  shows "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
-proof cases
-  assume "finite A" from this assms show ?thesis
-  proof induct
-    case (insert x A)
-    hence hyp: "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
-      by auto
-    { fix i assume *: "i \<in> insert x A"
-      hence "g i \<le> f i" using insert by simp
-      also have "f i < \<omega>" using * insert by (simp add: pextreal_less_\<omega>)
-      finally have "g i \<noteq> \<omega>" by (simp add: pextreal_less_\<omega>) }
-    hence "setsum g A \<noteq> \<omega>" "g x \<noteq> \<omega>" by (auto simp: setsum_\<omega>)
-    moreover have "setsum f A \<noteq> \<omega>" "f x \<noteq> \<omega>" using insert by (auto simp: setsum_\<omega>)
-    moreover have "g x \<le> f x" using insert by auto
-    moreover have "(\<Sum>i\<in>A. g i) \<le> (\<Sum>i\<in>A. f i)" using insert by (auto intro!: setsum_mono)
-    ultimately show ?case using `finite A` `x \<notin> A` hyp
-      by (auto simp: pextreal_noteq_omega_Ex)
-  qed simp
-qed simp
-
-lemma real_of_pextreal_diff:
-  "y \<le> x \<Longrightarrow> x \<noteq> \<omega> \<Longrightarrow> real x - real y = real (x - y)"
-  by (cases x, cases y) auto
-
-lemma psuminf_minus:
-  assumes ord: "\<And>i. g i \<le> f i" and fin: "psuminf g \<noteq> \<omega>" "psuminf f \<noteq> \<omega>"
-  shows "(\<Sum>\<^isub>\<infinity> i. f i - g i) = psuminf f - psuminf g"
-proof -
-  have [simp]: "\<And>i. f i \<noteq> \<omega>" using fin by (auto intro: psuminf_\<omega>)
-  from fin have "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity>x. f x)"
-    and "(\<lambda>x. real (g x)) sums real (\<Sum>\<^isub>\<infinity>x. g x)"
-    by (auto intro: psuminf_imp_suminf)
-  from sums_diff[OF this]
-  have "(\<lambda>n. real (f n - g n)) sums (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))" using fin ord
-    by (subst (asm) (1 2) real_of_pextreal_diff) (auto simp: psuminf_\<omega> psuminf_le)
-  hence "(\<Sum>\<^isub>\<infinity> i. Real (real (f i - g i))) = Real (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))"
-    by (rule suminf_imp_psuminf) simp
-  thus ?thesis using fin by (simp add: Real_real psuminf_\<omega>)
-qed
-
-lemma INF_eq_LIMSEQ:
-  assumes "mono (\<lambda>i. - f i)" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
-  shows "(INF n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
-proof
-  assume x: "(INF n. Real (f n)) = Real x"
-  { fix n
-    have "Real x \<le> Real (f n)" using x[symmetric] by (auto intro: INF_leI)
-    hence "x \<le> f n" using assms by simp
-    hence "\<bar>f n - x\<bar> = f n - x" by auto }
-  note abs_eq = this
-  show "f ----> x"
-  proof (rule LIMSEQ_I)
-    fix r :: real assume "0 < r"
-    show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
-    proof (rule ccontr)
-      assume *: "\<not> ?thesis"
-      { fix N
-        from * obtain n where *: "N \<le> n" "r \<le> f n - x"
-          using abs_eq by (auto simp: not_less)
-        hence "x + r \<le> f n" by auto
-        also have "f n \<le> f N" using `mono (\<lambda>i. - f i)` * by (auto dest: monoD)
-        finally have "Real (x + r) \<le> Real (f N)" using `0 \<le> f N` by auto }
-      hence "Real x < Real (x + r)"
-        and "Real (x + r) \<le> (INF n. Real (f n))" using `0 < r` `0 \<le> x` by (auto intro: le_INFI)
-      hence "Real x < (INF n. Real (f n))" by (rule less_le_trans)
-      thus False using x by auto
-    qed
-  qed
-next
-  assume "f ----> x"
-  show "(INF n. Real (f n)) = Real x"
-  proof (rule pextreal_INFI)
-    fix n
-    from decseq_le[OF _ `f ----> x`] assms
-    show "Real x \<le> Real (f n)" unfolding decseq_eq_incseq incseq_mono by auto
-  next
-    fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> y \<le> Real (f n)"
-    thus "y \<le> Real x"
-    proof (cases y)
-      case (preal r)
-      with * have "\<exists>N. \<forall>n\<ge>N. r \<le> f n" using assms by fastsimp
-      from LIMSEQ_le_const[OF `f ----> x` this]
-      show "y \<le> Real x" using `0 \<le> x` preal by auto
-    qed simp
-  qed
-qed
-
-lemma INFI_bound:
-  assumes "\<forall>N. x \<le> f N"
-  shows "x \<le> (INF n. f n)"
-  using assms by (simp add: INFI_def le_Inf_iff)
-
-lemma LIMSEQ_imp_lim_INF:
-  assumes pos: "\<And>i. 0 \<le> X i" and lim: "X ----> x"
-  shows "(SUP n. INF m. Real (X (n + m))) = Real x"
-proof -
-  have "0 \<le> x" using assms by (auto intro!: LIMSEQ_le_const)
-
-  have "\<And>n. (INF m. Real (X (n + m))) \<le> Real (X (n + 0))" by (rule INF_leI) simp
-  also have "\<And>n. Real (X (n + 0)) < \<omega>" by simp
-  finally have "\<forall>n. \<exists>r\<ge>0. (INF m. Real (X (n + m))) = Real r"
-    by (auto simp: pextreal_less_\<omega> pextreal_noteq_omega_Ex)
-  from choice[OF this] obtain r where r: "\<And>n. (INF m. Real (X (n + m))) = Real (r n)" "\<And>n. 0 \<le> r n"
-    by auto
-
-  show ?thesis unfolding r
-  proof (subst SUP_eq_LIMSEQ)
-    show "mono r" unfolding mono_def
-    proof safe
-      fix x y :: nat assume "x \<le> y"
-      have "Real (r x) \<le> Real (r y)" unfolding r(1)[symmetric] using pos
-      proof (safe intro!: INF_mono bexI)
-        fix m have "x + (m + y - x) = y + m"
-          using `x \<le> y` by auto
-        thus "Real (X (x + (m + y - x))) \<le> Real (X (y + m))" by simp
-      qed simp
-      thus "r x \<le> r y" using r by auto
-    qed
-    show "\<And>n. 0 \<le> r n" by fact
-    show "0 \<le> x" by fact
-    show "r ----> x"
-    proof (rule LIMSEQ_I)
-      fix e :: real assume "0 < e"
-      hence "0 < e/2" by auto
-      from LIMSEQ_D[OF lim this] obtain N where *: "\<And>n. N \<le> n \<Longrightarrow> \<bar>X n - x\<bar> < e/2"
-        by auto
-      show "\<exists>N. \<forall>n\<ge>N. norm (r n - x) < e"
-      proof (safe intro!: exI[of _ N])
-        fix n assume "N \<le> n"
-        show "norm (r n - x) < e"
-        proof cases
-          assume "r n < x"
-          have "x - r n \<le> e/2"
-          proof cases
-            assume e: "e/2 \<le> x"
-            have "Real (x - e/2) \<le> Real (r n)" unfolding r(1)[symmetric]
-            proof (rule le_INFI)
-              fix m show "Real (x - e / 2) \<le> Real (X (n + m))"
-                using *[of "n + m"] `N \<le> n`
-                using pos by (auto simp: field_simps abs_real_def split: split_if_asm)
-            qed
-            with e show ?thesis using pos `0 \<le> x` r(2) by auto
-          next
-            assume "\<not> e/2 \<le> x" hence "x - e/2 < 0" by auto
-            with `0 \<le> r n` show ?thesis by auto
-          qed
-          with `r n < x` show ?thesis by simp
-        next
-          assume e: "\<not> r n < x"
-          have "Real (r n) \<le> Real (X (n + 0))" unfolding r(1)[symmetric]
-            by (rule INF_leI) simp
-          hence "r n - x \<le> X n - x" using r pos by auto
-          also have "\<dots> < e/2" using *[OF `N \<le> n`] by (auto simp: field_simps abs_real_def split: split_if_asm)
-          finally have "r n - x < e" using `0 < e` by auto
-          with e show ?thesis by auto
-        qed
-      qed
-    qed
-  qed
-qed
-
-lemma real_of_pextreal_strict_mono_iff:
-  "real a < real b \<longleftrightarrow> (b \<noteq> \<omega> \<and> ((a = \<omega> \<and> 0 < b) \<or> (a < b)))"
-proof (cases a)
-  case infinite thus ?thesis by (cases b) auto
-next
-  case preal thus ?thesis by (cases b) auto
-qed
-
-lemma real_of_pextreal_mono_iff:
-  "real a \<le> real b \<longleftrightarrow> (a = \<omega> \<or> (b \<noteq> \<omega> \<and> a \<le> b) \<or> (b = \<omega> \<and> a = 0))"
-proof (cases a)
-  case infinite thus ?thesis by (cases b) auto
-next
-  case preal thus ?thesis by (cases b)  auto
-qed
-
-lemma ex_pextreal_inverse_of_nat_Suc_less:
-  fixes e :: pextreal assumes "0 < e" shows "\<exists>n. inverse (of_nat (Suc n)) < e"
-proof (cases e)
-  case (preal r)
-  with `0 < e` ex_inverse_of_nat_Suc_less[of r]
-  obtain n where "inverse (of_nat (Suc n)) < r" by auto
-  with preal show ?thesis
-    by (auto simp: real_eq_of_nat[symmetric])
-qed auto
-
-lemma Lim_eq_Sup_mono:
-  fixes u :: "nat \<Rightarrow> pextreal" assumes "mono u"
-  shows "u ----> (SUP i. u i)"
-proof -
-  from lim_pextreal_increasing[of u] `mono u`
-  obtain l where l: "u ----> l" unfolding mono_def by auto
-  from SUP_Lim_pextreal[OF _ this] `mono u`
-  have "(SUP i. u i) = l" unfolding mono_def by auto
-  with l show ?thesis by simp
-qed
-
-lemma isotone_Lim:
-  fixes x :: pextreal assumes "u \<up> x"
-  shows "u ----> x" (is ?lim) and "mono u" (is ?mono)
-proof -
-  show ?mono using assms unfolding mono_iff_le_Suc isoton_def by auto
-  from Lim_eq_Sup_mono[OF this] `u \<up> x`
-  show ?lim unfolding isoton_def by simp
-qed
-
-lemma isoton_iff_Lim_mono:
-  fixes u :: "nat \<Rightarrow> pextreal"
-  shows "u \<up> x \<longleftrightarrow> (mono u \<and> u ----> x)"
-proof safe
-  assume "mono u" and x: "u ----> x"
-  with SUP_Lim_pextreal[OF _ x]
-  show "u \<up> x" unfolding isoton_def
-    using `mono u`[unfolded mono_def]
-    using `mono u`[unfolded mono_iff_le_Suc]
-    by auto
-qed (auto dest: isotone_Lim)
-
-lemma pextreal_inverse_inverse[simp]:
-  fixes x :: pextreal
-  shows "inverse (inverse x) = x"
-  by (cases x) auto
-
-lemma atLeastAtMost_omega_eq_atLeast:
-  shows "{a .. \<omega>} = {a ..}"
-by auto
-
-lemma atLeast0AtMost_eq_atMost: "{0 :: pextreal .. a} = {.. a}" by auto
-
-lemma greaterThan_omega_Empty: "{\<omega> <..} = {}" by auto
-
-lemma lessThan_0_Empty: "{..< 0 :: pextreal} = {}" by auto
-
-lemma real_of_pextreal_inverse[simp]:
-  fixes X :: pextreal
-  shows "real (inverse X) = 1 / real X"
-  by (cases X) (auto simp: inverse_eq_divide)
-
-lemma real_of_pextreal_le_0[simp]: "real (X :: pextreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
-  by (cases X) auto
-
-lemma real_of_pextreal_less_0[simp]: "\<not> (real (X :: pextreal) < 0)"
-  by (cases X) auto
-
-lemma abs_real_of_pextreal[simp]: "\<bar>real (X :: pextreal)\<bar> = real X"
-  by simp
-
-lemma zero_less_real_of_pextreal: "0 < real (X :: pextreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
-  by (cases X) auto
-
-end
--- a/src/HOL/Probability/Probability.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Probability.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,6 +1,7 @@
 theory Probability
 imports
   Complete_Measure
+  Lebesgue_Measure
   Information
   "ex/Dining_Cryptographers"
   "ex/Koepf_Duermuth_Countermeasure"
--- a/src/HOL/Probability/Probability_Space.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Probability_Space.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -2,33 +2,33 @@
 imports Lebesgue_Integration Radon_Nikodym Product_Measure
 begin
 
-lemma real_of_pextreal_inverse[simp]:
-  fixes X :: pextreal
+lemma real_of_extreal_inverse[simp]:
+  fixes X :: extreal
   shows "real (inverse X) = 1 / real X"
   by (cases X) (auto simp: inverse_eq_divide)
 
-lemma real_of_pextreal_le_0[simp]: "real (X :: pextreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
+lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
+  by (cases X) auto
+
+lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
   by (cases X) auto
 
-lemma real_of_pextreal_less_0[simp]: "\<not> (real (X :: pextreal) < 0)"
+lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
   by (cases X) auto
 
+lemma real_of_extreal_le_1: fixes X :: extreal shows "X \<le> 1 \<Longrightarrow> real X \<le> 1"
+  by (cases X) (auto simp: one_extreal_def)
+
 locale prob_space = measure_space +
   assumes measure_space_1: "measure M (space M) = 1"
 
-lemma abs_real_of_pextreal[simp]: "\<bar>real (X :: pextreal)\<bar> = real X"
-  by simp
-
-lemma zero_less_real_of_pextreal: "0 < real (X :: pextreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
-  by (cases X) auto
-
 sublocale prob_space < finite_measure
 proof
-  from measure_space_1 show "\<mu> (space M) \<noteq> \<omega>" by simp
+  from measure_space_1 show "\<mu> (space M) \<noteq> \<infinity>" by simp
 qed
 
 abbreviation (in prob_space) "events \<equiv> sets M"
-abbreviation (in prob_space) "prob \<equiv> \<lambda>A. real (\<mu> A)"
+abbreviation (in prob_space) "prob \<equiv> \<mu>'"
 abbreviation (in prob_space) "prob_preserving \<equiv> measure_preserving"
 abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
 abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
@@ -40,53 +40,57 @@
   "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
 
 definition (in prob_space)
-  "distribution X = (\<lambda>s. \<mu> ((X -` s) \<inter> (space M)))"
+  "distribution X A = \<mu>' (X -` A \<inter> space M)"
 
 abbreviation (in prob_space)
   "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
 
+declare (in finite_measure) positive_measure'[intro, simp]
+
 lemma (in prob_space) distribution_cong:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
   shows "distribution X = distribution Y"
   unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>"])
+  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
 
 lemma (in prob_space) joint_distribution_cong:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
   assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
   shows "joint_distribution X Y = joint_distribution X' Y'"
   unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>"])
+  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
 
 lemma (in prob_space) distribution_id[simp]:
-  assumes "N \<in> sets M" shows "distribution (\<lambda>x. x) N = \<mu> N"
-  using assms by (auto simp: distribution_def intro!: arg_cong[where f=\<mu>])
+  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
+  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
 
 lemma (in prob_space) prob_space: "prob (space M) = 1"
-  unfolding measure_space_1 by simp
+  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
+
+lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
+  using bounded_measure[of A] by (simp add: prob_space)
+
+lemma (in prob_space) distribution_positive[simp, intro]:
+  "0 \<le> distribution X A" unfolding distribution_def by auto
 
-lemma (in prob_space) measure_le_1[simp, intro]:
-  assumes "A \<in> events" shows "\<mu> A \<le> 1"
-proof -
-  have "\<mu> A \<le> \<mu> (space M)"
-    using assms sets_into_space by(auto intro!: measure_mono)
-  also note measure_space_1
-  finally show ?thesis .
-qed
+lemma (in prob_space) joint_distribution_remove[simp]:
+    "joint_distribution X X {(x, x)} = distribution X {x}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+
+lemma (in prob_space) distribution_1:
+  "distribution X A \<le> 1"
+  unfolding distribution_def by simp
 
 lemma (in prob_space) prob_compl:
-  assumes "A \<in> events"
+  assumes A: "A \<in> events"
   shows "prob (space M - A) = 1 - prob A"
-  using `A \<in> events`[THEN sets_into_space] `A \<in> events` measure_space_1
-  by (subst real_finite_measure_Diff) auto
+  using finite_measure_compl[OF A] by (simp add: prob_space)
 
-lemma (in prob_space) indep_space:
-  assumes "s \<in> events"
-  shows "indep (space M) s"
-  using assms prob_space by (simp add: indep_def)
+lemma (in prob_space) indep_space: "s \<in> events \<Longrightarrow> indep (space M) s"
+  by (simp add: indep_def prob_space)
 
 lemma (in prob_space) prob_space_increasing: "increasing M prob"
-  by (auto intro!: real_measure_mono simp: increasing_def)
+  by (auto intro!: finite_measure_mono simp: increasing_def)
 
 lemma (in prob_space) prob_zero_union:
   assumes "s \<in> events" "t \<in> events" "prob t = 0"
@@ -94,9 +98,9 @@
 using assms
 proof -
   have "prob (s \<union> t) \<le> prob s"
-    using real_finite_measure_subadditive[of s t] assms by auto
+    using finite_measure_subadditive[of s t] assms by auto
   moreover have "prob (s \<union> t) \<ge> prob s"
-    using assms by (blast intro: real_measure_mono)
+    using assms by (blast intro: finite_measure_mono)
   ultimately show ?thesis by simp
 qed
 
@@ -127,9 +131,9 @@
 using assms
 proof -
   have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
-    by (rule real_finite_measure_UNION[OF assms(1,3)])
+    by (rule finite_measure_UNION[OF assms(1,3)])
   have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
-    by (rule real_finite_measure_UNION[OF assms(2,4)])
+    by (rule finite_measure_UNION[OF assms(2,4)])
   show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
 qed
 
@@ -139,10 +143,9 @@
   shows "prob (\<Union> i :: nat. c i) = 0"
 proof (rule antisym)
   show "prob (\<Union> i :: nat. c i) \<le> 0"
-    using real_finite_measure_countably_subadditive[OF assms(1)]
+    using finite_measure_countably_subadditive[OF assms(1)]
     by (simp add: assms(2) suminf_zero summable_zero)
-  show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pextreal_nonneg)
-qed
+qed simp
 
 lemma (in prob_space) indep_sym:
    "indep a b \<Longrightarrow> indep b a"
@@ -163,7 +166,7 @@
   from someI_ex[OF this] assms
   have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
   have "prob s = (\<Sum> x \<in> s. prob {x})"
-    using real_finite_measure_finite_singelton[OF s_finite] by simp
+    using finite_measure_finite_singleton[OF s_finite] by simp
   also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
   also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
     using setsum_constant assms by (simp add: real_eq_of_nat)
@@ -182,7 +185,7 @@
     using `e \<in> events` sets_into_space upper by blast
   hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
   also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
-  proof (rule real_finite_measure_finite_Union)
+  proof (rule finite_measure_finite_Union)
     show "finite s" by fact
     show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
     show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
@@ -193,64 +196,62 @@
 
 lemma (in prob_space) distribution_prob_space:
   assumes "random_variable S X"
-  shows "prob_space (S\<lparr>measure := distribution X\<rparr>)"
+  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
 proof -
-  interpret S: measure_space "S\<lparr>measure := distribution X\<rparr>"
-    unfolding distribution_def using assms
-    by (intro measure_space_vimage)
-       (auto intro!: sigma_algebra.sigma_algebra_cong[of S] simp: measure_preserving_def)
+  interpret S: measure_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
+  proof (rule measure_space.measure_space_cong)
+    show "measure_space (S\<lparr> measure := \<lambda>A. \<mu> (X -` A \<inter> space M) \<rparr>)"
+      using assms by (auto intro!: measure_space_vimage simp: measure_preserving_def)
+  qed (insert assms, auto simp add: finite_measure_eq distribution_def measurable_sets)
   show ?thesis
   proof (default, simp)
     have "X -` space S \<inter> space M = space M"
       using `random_variable S X` by (auto simp: measurable_def)
-    then show "distribution X (space S) = 1"
-      using measure_space_1 by (simp add: distribution_def)
+    then show "extreal (distribution X (space S)) = 1"
+      by (simp add: distribution_def one_extreal_def prob_space)
   qed
 qed
 
 lemma (in prob_space) AE_distribution:
-  assumes X: "random_variable MX X" and "measure_space.almost_everywhere (MX\<lparr>measure := distribution X\<rparr>) (\<lambda>x. Q x)"
+  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
   shows "AE x. Q (X x)"
 proof -
-  interpret X: prob_space "MX\<lparr>measure := distribution X\<rparr>" using X by (rule distribution_prob_space)
+  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
   obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
     using assms unfolding X.almost_everywhere_def by auto
-  show "AE x. Q (X x)"
-    using X[unfolded measurable_def] N unfolding distribution_def
-    by (intro AE_I'[where N="X -` N \<inter> space M"]) auto
+  from X[unfolded measurable_def] N show "AE x. Q (X x)"
+    by (intro AE_I'[where N="X -` N \<inter> space M"])
+       (auto simp: finite_measure_eq distribution_def measurable_sets)
 qed
 
-lemma (in prob_space) distribution_lebesgue_thm1:
-  assumes "random_variable s X"
-  assumes "A \<in> sets s"
-  shows "real (distribution X A) = expectation (indicator (X -` A \<inter> space M))"
-unfolding distribution_def
-using assms unfolding measurable_def
-using integral_indicator by auto
+lemma (in prob_space) distribution_eq_integral:
+  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
+  using finite_measure_eq[of "X -` A \<inter> space M"]
+  by (auto simp: measurable_sets distribution_def)
 
-lemma (in prob_space) distribution_lebesgue_thm2:
-  assumes "random_variable S X" and "A \<in> sets S"
-  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := distribution X\<rparr>) (indicator A)"
+lemma (in prob_space) distribution_eq_translated_integral:
+  assumes "random_variable S X" "A \<in> sets S"
+  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
 proof -
-  interpret S: prob_space "S\<lparr>measure := distribution X\<rparr>"
+  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
     using assms(1) by (rule distribution_prob_space)
   show ?thesis
-    using S.positive_integral_indicator(1)
-    using assms unfolding distribution_def by auto
+    using S.positive_integral_indicator(1)[of A] assms by simp
 qed
 
 lemma (in prob_space) finite_expectation1:
   assumes f: "finite (X`space M)" and rv: "random_variable borel X"
-  shows "expectation X = (\<Sum> r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))"
-proof (rule integral_on_finite(2)[OF rv[THEN conjunct2] f])
-  fix x have "X -` {x} \<inter> space M \<in> sets M"
-    using rv unfolding measurable_def by auto
-  thus "\<mu> (X -` {x} \<inter> space M) \<noteq> \<omega>" using finite_measure by simp
+  shows "expectation X = (\<Sum>r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))" (is "_ = ?r")
+proof (subst integral_on_finite)
+  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
+  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
+    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
+    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
 qed
 
 lemma (in prob_space) finite_expectation:
   assumes "finite (X`space M)" "random_variable borel X"
-  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))"
+  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
   using assms unfolding distribution_def using finite_expectation1 by auto
 
 lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
@@ -267,23 +268,18 @@
 lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
 proof -
   have "X -` X ` space M \<inter> space M = space M" by auto
-  thus ?thesis unfolding distribution_def by (simp add: measure_space_1)
+  thus ?thesis unfolding distribution_def by (simp add: prob_space)
 qed
 
 lemma (in prob_space) distribution_one:
   assumes "random_variable M' X" and "A \<in> sets M'"
   shows "distribution X A \<le> 1"
 proof -
-  have "distribution X A \<le> \<mu> (space M)" unfolding distribution_def
-    using assms[unfolded measurable_def] by (auto intro!: measure_mono)
-  thus ?thesis by (simp add: measure_space_1)
+  have "distribution X A \<le> \<mu>' (space M)" unfolding distribution_def
+    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
+  thus ?thesis by (simp add: prob_space)
 qed
 
-lemma (in prob_space) distribution_finite:
-  assumes "random_variable M' X" and "A \<in> sets M'"
-  shows "distribution X A \<noteq> \<omega>"
-  using distribution_one[OF assms] by auto
-
 lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
   assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
     (is "random_variable ?S X")
@@ -302,8 +298,8 @@
   have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
     "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
     using `y \<noteq> x` by auto
-  with measure_inter_full_set[OF single single, of x y] assms(2)
-  show ?thesis unfolding distribution_def measure_space_1 by auto
+  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
+  show ?thesis by (auto simp: distribution_def prob_space)
 next
   assume "{y} \<notin> sets ?S"
   then have "X -` {y} \<inter> space M = {}" by auto
@@ -315,20 +311,17 @@
     and A: "A \<in> sets MX" and B: "B \<in> sets MY"
   shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
   unfolding distribution_def
-proof (intro measure_mono)
+proof (intro finite_measure_mono)
   show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
   show "X -` A \<inter> space M \<in> events"
     using X A unfolding measurable_def by simp
   have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
     (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
-  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<in> events"
-    unfolding * apply (rule Int)
-    using assms unfolding measurable_def by auto
 qed
 
 lemma (in prob_space) joint_distribution_commute:
   "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
 
 lemma (in prob_space) joint_distribution_Times_le_snd:
   assumes X: "random_variable MX X" and Y: "random_variable MY Y"
@@ -352,27 +345,14 @@
     unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
 qed
 
-lemma (in prob_space) distribution_order:
-  assumes "random_variable MX X" "random_variable MY Y"
-  assumes "{x} \<in> sets MX" "{y} \<in> sets MY"
-  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
-    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
-    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
-    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
-    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-  using joint_distribution_Times_le_snd[OF assms]
-  using joint_distribution_Times_le_fst[OF assms]
-  by auto
-
 lemma (in prob_space) joint_distribution_commute_singleton:
   "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
 
 lemma (in prob_space) joint_distribution_assoc_singleton:
   "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
    joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
 
 locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2
 
@@ -383,13 +363,13 @@
 
 lemma countably_additiveI[case_names countably]:
   assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
-    (\<Sum>\<^isub>\<infinity>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+    (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
   shows "countably_additive M \<mu>"
   using assms unfolding countably_additive_def by auto
 
 lemma (in prob_space) joint_distribution_prob_space:
   assumes "random_variable MX X" "random_variable MY Y"
-  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := joint_distribution X Y\<rparr>)"
+  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
   using random_variable_pairI[OF assms] by (rule distribution_prob_space)
 
 section "Probability spaces on finite sets"
@@ -407,20 +387,13 @@
 
 lemma (in prob_space) distribution_finite_prob_space:
   assumes "finite_random_variable MX X"
-  shows "finite_prob_space (MX\<lparr>measure := distribution X\<rparr>)"
+  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
 proof -
-  interpret X: prob_space "MX\<lparr>measure := distribution X\<rparr>"
+  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
     using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
   interpret MX: finite_sigma_algebra MX
     using assms by auto
-  show ?thesis
-  proof (default, simp_all)
-    fix x assume "x \<in> space MX"
-    then have "X -` {x} \<inter> space M \<in> sets M"
-      using assms unfolding measurable_def by simp
-    then show "distribution X {x} \<noteq> \<omega>"
-      unfolding distribution_def by simp
-  qed (rule MX.finite_space)
+  show ?thesis by default (simp_all add: MX.finite_space)
 qed
 
 lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
@@ -451,9 +424,9 @@
   by (auto dest!: finite_random_variableD)
 
 lemma (in prob_space) sum_over_space_real_distribution:
-  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
   unfolding distribution_def prob_space[symmetric]
-  by (subst real_finite_measure_finite_Union[symmetric])
+  by (subst finite_measure_finite_Union[symmetric])
      (auto simp add: disjoint_family_on_def simple_function_def
            intro!: arg_cong[where f=prob])
 
@@ -475,7 +448,7 @@
   "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
   unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
 
-lemma (in prob_space) finite_random_variable_vimage:
+lemma (in prob_space) finite_random_variable_measurable:
   assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
 proof -
   interpret X: finite_sigma_algebra MX using X by simp
@@ -492,15 +465,12 @@
   assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
   shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
   unfolding distribution_def
-proof (intro measure_mono)
+proof (intro finite_measure_mono)
   show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
   show "X -` A \<inter> space M \<in> events"
-    using finite_random_variable_vimage[OF X] .
+    using finite_random_variable_measurable[OF X] .
   have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
     (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
-  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<in> events"
-    unfolding * apply (rule Int)
-    using assms[THEN finite_random_variable_vimage] by auto
 qed
 
 lemma (in prob_space) joint_distribution_finite_Times_le_snd:
@@ -511,6 +481,7 @@
      (simp add: swap_product joint_distribution_finite_Times_le_fst)
 
 lemma (in prob_space) finite_distribution_order:
+  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
   assumes "finite_random_variable MX X" "finite_random_variable MY Y"
   shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
     and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
@@ -520,26 +491,14 @@
     and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
   using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
   using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
-  by auto
-
-lemma (in prob_space) finite_distribution_finite:
-  assumes "finite_random_variable M' X"
-  shows "distribution X {x} \<noteq> \<omega>"
-proof -
-  have "distribution X {x} \<le> \<mu> (space M)"
-    unfolding distribution_def
-    using finite_random_variable_vimage[OF assms]
-    by (intro measure_mono) auto
-  then show ?thesis
-    by auto
-qed
+  by (auto intro: antisym)
 
 lemma (in prob_space) setsum_joint_distribution:
   assumes X: "finite_random_variable MX X"
   assumes Y: "random_variable MY Y" "B \<in> sets MY"
   shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
   unfolding distribution_def
-proof (subst measure_finitely_additive'')
+proof (subst finite_measure_finite_Union[symmetric])
   interpret MX: finite_sigma_algebra MX using X by auto
   show "finite (space MX)" using MX.finite_space .
   let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
@@ -549,9 +508,8 @@
       using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
       using MX.sets_eq_Pow by auto }
   show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
-  show "\<mu> (\<Union>i\<in>space MX. ?d i) = \<mu> (Y -` B \<inter> space M)"
-    using X[unfolded measurable_def]
-    by (auto intro!: arg_cong[where f=\<mu>])
+  show "\<mu>' (\<Union>i\<in>space MX. ?d i) = \<mu>' (Y -` B \<inter> space M)"
+    using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\<mu>'])
 qed
 
 lemma (in prob_space) setsum_joint_distribution_singleton:
@@ -562,28 +520,6 @@
     finite_random_variableD[OF Y(1)]
     finite_random_variable_imp_sets[OF Y]] by simp
 
-lemma (in prob_space) setsum_real_joint_distribution:
-  fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme"
-  assumes X: "finite_random_variable MX X"
-  assumes Y: "random_variable MY Y" "B \<in> sets MY"
-  shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y ({a} \<times> B))) = real (distribution Y B)"
-proof -
-  interpret MX: finite_sigma_algebra MX using X by auto
-  show ?thesis
-    unfolding setsum_joint_distribution[OF assms, symmetric]
-    using distribution_finite[OF random_variable_pairI[OF finite_random_variableD[OF X] Y(1)]] Y(2)
-    by (simp add: space_pair_measure real_of_pextreal_setsum)
-qed
-
-lemma (in prob_space) setsum_real_joint_distribution_singleton:
-  fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme"
-  assumes X: "finite_random_variable MX X"
-  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
-  shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y {(a,b)})) = real (distribution Y {b})"
-  using setsum_real_joint_distribution[OF X
-    finite_random_variableD[OF Y(1)]
-    finite_random_variable_imp_sets[OF Y]] by simp
-
 locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
 
 sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 M2 by default
@@ -593,7 +529,7 @@
 lemma (in prob_space) joint_distribution_finite_prob_space:
   assumes X: "finite_random_variable MX X"
   assumes Y: "finite_random_variable MY Y"
-  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := joint_distribution X Y\<rparr>)"
+  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
   by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
 
 lemma finite_prob_space_eq:
@@ -602,18 +538,15 @@
   by auto
 
 lemma (in prob_space) not_empty: "space M \<noteq> {}"
-  using prob_space empty_measure by auto
+  using prob_space empty_measure' by auto
 
 lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
   using measure_space_1 sum_over_space by simp
 
-lemma (in finite_prob_space) positive_distribution: "0 \<le> distribution X x"
-  unfolding distribution_def by simp
-
 lemma (in finite_prob_space) joint_distribution_restriction_fst:
   "joint_distribution X Y A \<le> distribution X (fst ` A)"
   unfolding distribution_def
-proof (safe intro!: measure_mono)
+proof (safe intro!: finite_measure_mono)
   fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
   show "x \<in> X -` fst ` A"
     by (auto intro!: image_eqI[OF _ *])
@@ -622,7 +555,7 @@
 lemma (in finite_prob_space) joint_distribution_restriction_snd:
   "joint_distribution X Y A \<le> distribution Y (snd ` A)"
   unfolding distribution_def
-proof (safe intro!: measure_mono)
+proof (safe intro!: finite_measure_mono)
   fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
   show "x \<in> Y -` snd ` A"
     by (auto intro!: image_eqI[OF _ *])
@@ -637,17 +570,16 @@
   and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
   and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
   and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
-  using positive_distribution[of X x']
-    positive_distribution[of "\<lambda>x. (X x, Y x)" "{(x, y)}"]
+  using
     joint_distribution_restriction_fst[of X Y "{(x, y)}"]
     joint_distribution_restriction_snd[of X Y "{(x, y)}"]
-  by auto
+  by (auto intro: antisym)
 
 lemma (in finite_prob_space) distribution_mono:
   assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
   shows "distribution X x \<le> distribution Y y"
   unfolding distribution_def
-  using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
+  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
 
 lemma (in finite_prob_space) distribution_mono_gt_0:
   assumes gt_0: "0 < distribution X x"
@@ -657,49 +589,21 @@
 
 lemma (in finite_prob_space) sum_over_space_distrib:
   "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def measure_space_1[symmetric] using finite_space
-  by (subst measure_finitely_additive'')
-     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>])
+  unfolding distribution_def prob_space[symmetric] using finite_space
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow
+           intro!: arg_cong[where f=\<mu>'])
 
 lemma (in finite_prob_space) sum_over_space_real_distribution:
-  "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
   unfolding distribution_def prob_space[symmetric] using finite_space
-  by (subst real_finite_measure_finite_Union[symmetric])
+  by (subst finite_measure_finite_Union[symmetric])
      (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
 
 lemma (in finite_prob_space) finite_sum_over_space_eq_1:
-  "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1"
-  using sum_over_space_eq_1 finite_measure by (simp add: real_of_pextreal_setsum sets_eq_Pow)
-
-lemma (in finite_prob_space) distribution_finite:
-  "distribution X A \<noteq> \<omega>"
-  using finite_measure[of "X -` A \<inter> space M"]
-  unfolding distribution_def sets_eq_Pow by auto
-
-lemma (in finite_prob_space) real_distribution_gt_0[simp]:
-  "0 < real (distribution Y y) \<longleftrightarrow>  0 < distribution Y y"
-  using assms by (auto intro!: real_pextreal_pos distribution_finite)
-
-lemma (in finite_prob_space) real_distribution_mult_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y * distribution X x)"
-  unfolding real_of_pextreal_mult[symmetric]
-  using assms by (auto intro!: mult_pos_pos)
-
-lemma (in finite_prob_space) real_distribution_divide_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y / distribution X x)"
-  unfolding divide_pextreal_def real_of_pextreal_mult[symmetric]
-  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
-
-lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y * inverse (distribution X x))"
-  unfolding divide_pextreal_def real_of_pextreal_mult[symmetric]
-  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
+  "(\<Sum>x\<in>space M. prob {x}) = 1"
+  using prob_space finite_space
+  by (subst (asm) finite_measure_finite_singleton) auto
 
 lemma (in prob_space) distribution_remove_const:
   shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
@@ -707,8 +611,7 @@
   and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
   and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
   and "distribution (\<lambda>x. ()) {()} = 1"
-  unfolding measure_space_1[symmetric]
-  by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def)
+  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
 
 lemma (in finite_prob_space) setsum_distribution_gen:
   assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
@@ -716,7 +619,7 @@
   shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
   unfolding distribution_def assms
   using finite_space assms
-  by (subst measure_finitely_additive'')
+  by (subst finite_measure_finite_Union[symmetric])
      (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
       intro!: arg_cong[where f=prob])
 
@@ -728,61 +631,17 @@
   "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
   by (auto intro!: inj_onI setsum_distribution_gen)
 
-lemma (in finite_prob_space) setsum_real_distribution_gen:
-  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
-  and "inj_on f (X`space M)"
-  shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})"
-  unfolding distribution_def assms
-  using finite_space assms
-  by (subst real_finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
-        intro!: arg_cong[where f=prob])
-
-lemma (in finite_prob_space) setsum_real_distribution:
-  "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})"
-  "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})"
-  "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})"
-  "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})"
-  "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})"
-  by (auto intro!: inj_onI setsum_real_distribution_gen)
-
-lemma (in finite_prob_space) real_distribution_order:
-  shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})"
-  and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})"
-  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})"
-  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})"
-  and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
-  using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
-  using real_pextreal_nonneg[of "joint_distribution X Y {(x, y)}"]
-  by auto
-
-lemma (in prob_space) joint_distribution_remove[simp]:
-    "joint_distribution X X {(x, x)} = distribution X {x}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"])
-
-lemma (in finite_prob_space) distribution_1:
-  "distribution X A \<le> 1"
-  unfolding distribution_def measure_space_1[symmetric]
-  by (auto intro!: measure_mono simp: sets_eq_Pow)
-
-lemma (in finite_prob_space) real_distribution_1:
-  "real (distribution X A) \<le> 1"
-  unfolding real_pextreal_1[symmetric]
-  by (rule real_of_pextreal_mono[OF _ distribution_1]) simp
-
 lemma (in finite_prob_space) uniform_prob:
   assumes "x \<in> space M"
   assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
-  shows "prob {x} = 1 / real (card (space M))"
+  shows "prob {x} = 1 / card (space M)"
 proof -
   have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
     using assms(2)[OF _ `x \<in> space M`] by blast
   have "1 = prob (space M)"
     using prob_space by auto
   also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
-    using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
+    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
       sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
       finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
     by (auto simp add:setsum_restrict_set)
@@ -809,7 +668,7 @@
 qed
 
 lemma (in prob_space) prob_space_of_restricted_space:
-  assumes "\<mu> A \<noteq> 0" "\<mu> A \<noteq> \<omega>" "A \<in> sets M"
+  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
   shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
     (is "prob_space ?P")
 proof -
@@ -820,29 +679,48 @@
   show "prob_space ?P"
   proof
     show "measure ?P (space ?P) = 1"
-      using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pextreal_noteq_omega_Ex)
-    show "measure ?P {} = 0" by auto
-    have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
-    then show "countably_additive ?P (measure ?P)"
-        unfolding countably_additive_def psuminf_cmult_right
-        using A.measure_countably_additive by auto
+      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
+    show "positive ?P (measure ?P)"
+    proof (simp add: positive_def, safe)
+      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
+      fix B assume "B \<in> events"
+      with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
+      show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
+    qed
+    show "countably_additive ?P (measure ?P)"
+    proof (simp add: countably_additive_def, safe)
+      fix B and F :: "nat \<Rightarrow> 'a set"
+      assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
+      { fix i
+        from F have "F i \<in> op \<inter> A ` events" by auto
+        with `A \<in> events` have "F i \<in> events" by auto }
+      moreover then have "range F \<subseteq> events" by auto
+      moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
+        by (simp add: mult_commute divide_extreal_def)
+      moreover have "0 \<le> inverse (\<mu> A)"
+        using real_measure[OF `A \<in> events`] by auto
+      ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
+        using measure_countably_additive[of F] F
+        by (auto simp: suminf_cmult_extreal)
+    qed
   qed
 qed
 
 lemma finite_prob_spaceI:
-  assumes "finite (space M)" "sets M = Pow(space M)" "measure M (space M) = 1" "measure M {} = 0"
+  assumes "finite (space M)" "sets M = Pow(space M)"
+    and "measure M (space M) = 1" "measure M {} = 0" "\<And>A. A \<subseteq> space M \<Longrightarrow> 0 \<le> measure M A"
     and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M B"
   shows "finite_prob_space M"
   unfolding finite_prob_space_eq
 proof
   show "finite_measure_space M" using assms
-     by (auto intro!: finite_measure_spaceI)
+    by (auto intro!: finite_measure_spaceI)
   show "measure M (space M) = 1" by fact
 qed
 
 lemma (in finite_prob_space) finite_measure_space:
   fixes X :: "'a \<Rightarrow> 'x"
-  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
+  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
     (is "finite_measure_space ?S")
 proof (rule finite_measure_spaceI, simp_all)
   show "finite (X ` space M)" using finite_space by simp
@@ -850,51 +728,41 @@
   fix A B :: "'x set" assume "A \<inter> B = {}"
   then show "distribution X (A \<union> B) = distribution X A + distribution X B"
     unfolding distribution_def
-    by (subst measure_additive)
-       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
+    by (subst finite_measure_Union[symmetric])
+       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
 qed
 
 lemma (in finite_prob_space) finite_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X \<rparr>"
-  by (simp add: finite_prob_space_eq finite_measure_space)
-
-lemma (in finite_prob_space) real_distribution_order':
-  shows "real (distribution X {x}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  and "real (distribution Y {y}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
-  using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
-  using real_pextreal_nonneg[of "joint_distribution X Y {(x, y)}"]
-  by auto
+  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
+  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
 
 lemma (in finite_prob_space) finite_product_measure_space:
   fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
   assumes "finite s1" "finite s2"
-  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = joint_distribution X Y\<rparr>"
+  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
     (is "finite_measure_space ?M")
 proof (rule finite_measure_spaceI, simp_all)
   show "finite (s1 \<times> s2)"
     using assms by auto
-  show "joint_distribution X Y (s1\<times>s2) \<noteq> \<omega>"
-    using distribution_finite .
 next
   fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
   then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
     unfolding distribution_def
-    by (subst measure_additive)
-       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
+    by (subst finite_measure_Union[symmetric])
+       (auto intro!: arg_cong[where f=\<mu>'] simp: sets_eq_Pow)
 qed
 
 lemma (in finite_prob_space) finite_product_measure_space_of_images:
   shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
                                 sets = Pow (X ` space M \<times> Y ` space M),
-                                measure = joint_distribution X Y \<rparr>"
+                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
   using finite_space by (auto intro!: finite_product_measure_space)
 
 lemma (in finite_prob_space) finite_product_prob_space_of_images:
   "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
-                       measure = joint_distribution X Y \<rparr>"
+                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
   (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
   have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
   thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
     by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
@@ -903,11 +771,11 @@
 section "Conditional Expectation and Probability"
 
 lemma (in prob_space) conditional_expectation_exists:
-  fixes X :: "'a \<Rightarrow> pextreal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M"
+  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
-  shows "\<exists>Y\<in>borel_measurable N. \<forall>C\<in>sets N.
-      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
+  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
+      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
 proof -
   note N(4)[simp]
   interpret P: prob_space N
@@ -926,34 +794,29 @@
     unfolding P.absolutely_continuous_def
   proof safe
     fix A assume "A \<in> sets N" "P.\<mu> A = 0"
-    moreover then have f_borel: "?f A \<in> borel_measurable M"
-      using borel N by (auto intro: borel_measurable_indicator)
-    moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A"
-      by (auto simp: indicator_def)
-    moreover have "\<mu> \<dots> \<le> \<mu> A"
-      using `A \<in> sets N` N f_borel
-      by (auto intro!: measure_mono Int[of _ A] measurable_sets)
-    ultimately show "?Q A = 0"
-      by (simp add: positive_integral_0_iff)
+    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
+      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
+    then show "?Q A = 0"
+      by (auto simp add: positive_integral_0_iff_AE)
   qed
   from P.Radon_Nikodym[OF Q this]
-  obtain Y where Y: "Y \<in> borel_measurable N"
+  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
     "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
     by blast
   with N(2) show ?thesis
-    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,4,1)])
+    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
 qed
 
 definition (in prob_space)
-  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N
+  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
     \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
 
 abbreviation (in prob_space)
   "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
 
 lemma (in prob_space)
-  fixes X :: "'a \<Rightarrow> pextreal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M"
+  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
   shows borel_measurable_conditional_expectation:
     "conditional_expectation N X \<in> borel_measurable N"
@@ -970,36 +833,23 @@
     unfolding conditional_expectation_def by (rule someI2_ex) blast
 qed
 
-lemma (in sigma_algebra) factorize_measurable_function:
-  fixes Z :: "'a \<Rightarrow> pextreal" and Y :: "'a \<Rightarrow> 'c"
+lemma (in sigma_algebra) factorize_measurable_function_pos:
+  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
   assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
-  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
-    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
-proof safe
+  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
+  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
+proof -
   interpret M': sigma_algebra M' by fact
   have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
   from M'.sigma_algebra_vimage[OF this]
   interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 
-  { fix g :: "'c \<Rightarrow> pextreal" assume "g \<in> borel_measurable M'"
-    with M'.measurable_vimage_algebra[OF Y]
-    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by (rule measurable_comp)
-    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
-    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
-       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-       by (auto intro!: measurable_cong)
-    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by simp }
-
-  assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-  from va.borel_measurable_implies_simple_function_sequence[OF this]
-  obtain f where f: "\<And>i. simple_function (M'.vimage_algebra (space M) Y) (f i)" and "f \<up> Z" by blast
+  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
 
   have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
   proof
     fix i
-    from f[of i] have "finite (f i`space M)" and B_ex:
+    from f(1)[of i] have "finite (f i`space M)" and B_ex:
       "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
       unfolding simple_function_def by auto
     from B_ex[THEN bchoice] guess B .. note B = this
@@ -1011,23 +861,67 @@
       show "simple_function M' ?g" using B by auto
 
       fix x assume "x \<in> space M"
-      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pextreal)"
+      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
         unfolding indicator_def using B by auto
-      then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i]
+      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
         by (subst va.simple_function_indicator_representation) auto
     qed
   qed
   from choice[OF this] guess g .. note g = this
 
-  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
+  show ?thesis
   proof (intro ballI bexI)
     show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
       using g by (auto intro: M'.borel_measurable_simple_function)
     fix x assume "x \<in> space M"
-    have "Z x = (SUP i. f i) x" using `f \<up> Z` unfolding isoton_def by simp
-    also have "\<dots> = (SUP i. g i (Y x))" unfolding SUPR_apply
+    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
+    also have "\<dots> = (SUP i. g i (Y x))"
       using g `x \<in> space M` by simp
-    finally show "Z x = (SUP i. g i (Y x))" .
+    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
+  qed
+qed
+
+lemma extreal_0_le_iff_le_0[simp]:
+  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
+  by (cases rule: extreal2_cases[of a]) auto
+
+lemma (in sigma_algebra) factorize_measurable_function:
+  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
+  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
+  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
+    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
+proof safe
+  interpret M': sigma_algebra M' by fact
+  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
+  from M'.sigma_algebra_vimage[OF this]
+  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
+
+  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
+    with M'.measurable_vimage_algebra[OF Y]
+    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by (rule measurable_comp)
+    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
+    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
+       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+       by (auto intro!: measurable_cong)
+    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by simp }
+
+  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
+    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+    by auto
+  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
+  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
+  let "?g x" = "p x - n x"
+  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
+  proof (intro bexI ballI)
+    show "?g \<in> borel_measurable M'" using p n by auto
+    fix x assume "x \<in> space M"
+    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
+      using p n by auto
+    then show "Z x = ?g (Y x)"
+      by (auto split: split_max)
   qed
 qed
 
--- a/src/HOL/Probability/Product_Measure.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Product_Measure.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -416,7 +416,7 @@
   show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
 qed
 
-lemma (in pair_sigma_algebra) measurable_cut_fst:
+lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]:
   assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
 proof -
   let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
@@ -504,15 +504,14 @@
     fix A assume "A \<in> sets ?D"
     with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
         (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
-      by (auto intro!: M2.finite_measure_compl measurable_cut_fst
-               simp: vimage_Diff)
+      by (auto intro!: M2.measure_compl simp: vimage_Diff)
     with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
-      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pextreal_diff)
+      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff)
   next
     fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
-    moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>\<^isub>\<infinity> i. ?s (F i) x)"
+    moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
       by (intro M2.measure_countably_additive[symmetric])
-         (auto intro!: measurable_cut_fst simp: disjoint_family_on_def)
+         (auto simp: disjoint_family_on_def)
     ultimately show "(\<Union>i. F i) \<in> sets ?D"
       by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
   qed
@@ -546,7 +545,7 @@
   have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
   have M1: "sigma_finite_measure M1" by default
   from M2.disjoint_sigma_finite guess F .. note F = this
-  then have "\<And>i. F i \<in> sets M2" by auto
+  then have F_sets: "\<And>i. F i \<in> sets M2" by auto
   let "?C x i" = "F i \<inter> Pair x -` Q"
   { fix i
     let ?R = "M2.restricted_space (F i)"
@@ -578,10 +577,10 @@
       by simp }
   moreover
   { fix x
-    have "(\<Sum>\<^isub>\<infinity>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
+    have "(\<Sum>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
     proof (intro M2.measure_countably_additive)
       show "range (?C x) \<subseteq> sets M2"
-        using F `Q \<in> sets P` by (auto intro!: M2.Int measurable_cut_fst)
+        using F `Q \<in> sets P` by (auto intro!: M2.Int)
       have "disjoint_family F" using F by auto
       show "disjoint_family (?C x)"
         by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
@@ -589,10 +588,10 @@
     also have "(\<Union>i. ?C x i) = Pair x -` Q"
       using F sets_into_space `Q \<in> sets P`
       by (auto simp: space_pair_measure)
-    finally have "measure M2 (Pair x -` Q) = (\<Sum>\<^isub>\<infinity>i. measure M2 (?C x i))"
+    finally have "measure M2 (Pair x -` Q) = (\<Sum>i. measure M2 (?C x i))"
       by simp }
-  ultimately show ?thesis
-    by (auto intro!: M1.borel_measurable_psuminf)
+  ultimately show ?thesis using `Q \<in> sets P` F_sets
+    by (auto intro!: M1.borel_measurable_psuminf M2.Int)
 qed
 
 lemma (in pair_sigma_finite) measure_cut_measurable_snd:
@@ -620,7 +619,7 @@
   apply (simp add: pair_measure_def pair_measure_generator_def)
 proof (rule M1.positive_integral_cong)
   fix x assume "x \<in> space M1"
-  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: pextreal)"
+  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)"
     unfolding indicator_def by auto
   show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
     unfolding *
@@ -639,18 +638,21 @@
     by (simp add: M1.positive_integral_cmult_indicator ac_simps)
 qed
 
+lemma (in measure_space) measure_not_negative[simp,intro]:
+  assumes A: "A \<in> sets M" shows "\<mu> A \<noteq> - \<infinity>"
+  using positive_measure[OF A] by auto
+
 lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
-  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> F \<up> space E \<and>
-    (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<omega>)"
+  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> incseq F \<and> (\<Union>i. F i) = space E \<and>
+    (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
 proof -
   obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
-    F1: "range F1 \<subseteq> sets M1" "F1 \<up> space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<omega>" and
-    F2: "range F2 \<subseteq> sets M2" "F2 \<up> space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<omega>"
+    F1: "range F1 \<subseteq> sets M1" "incseq F1" "(\<Union>i. F1 i) = space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<infinity>" and
+    F2: "range F2 \<subseteq> sets M2" "incseq F2" "(\<Union>i. F2 i) = space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<infinity>"
     using M1.sigma_finite_up M2.sigma_finite_up by auto
-  then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)"
-    unfolding isoton_def by auto
+  then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
   let ?F = "\<lambda>i. F1 i \<times> F2 i"
-  show ?thesis unfolding isoton_def space_pair_measure
+  show ?thesis unfolding space_pair_measure
   proof (intro exI[of _ ?F] conjI allI)
     show "range ?F \<subseteq> sets E" using F1 F2
       by (fastsimp intro!: pair_measure_generatorI)
@@ -661,8 +663,8 @@
       then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
         by (auto simp: space)
       then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
-        using `F1 \<up> space M1` `F2 \<up> space M2`
-        by (auto simp: max_def dest: isoton_mono_le)
+        using `incseq F1` `incseq F2` unfolding incseq_def
+        by (force split: split_max)+
       then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
         by (intro SigmaI) (auto simp add: min_max.sup_commute)
       then show "x \<in> (\<Union>i. ?F i)" by auto
@@ -670,21 +672,22 @@
     then show "(\<Union>i. ?F i) = space E"
       using space by (auto simp: space pair_measure_generator_def)
   next
-    fix i show "F1 i \<times> F2 i \<subseteq> F1 (Suc i) \<times> F2 (Suc i)"
-      using `F1 \<up> space M1` `F2 \<up> space M2` unfolding isoton_def
-      by auto
+    fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
+      using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
   next
     fix i
     from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
-    with F1 F2 show "measure P (F1 i \<times> F2 i) \<noteq> \<omega>"
+    with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)]
+    show "measure P (F1 i \<times> F2 i) \<noteq> \<infinity>"
       by (simp add: pair_measure_times)
   qed
 qed
 
 sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
 proof
-  show "measure P {} = 0"
-    unfolding pair_measure_def pair_measure_generator_def sigma_def by auto
+  show "positive P (measure P)"
+    unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def
+    by (auto intro: M1.positive_integral_positive M2.positive_integral_positive)
 
   show "countably_additive P (measure P)"
     unfolding countably_additive_def
@@ -697,20 +700,20 @@
     moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
       by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
     moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
-      using F by (auto intro!: measurable_cut_fst)
-    ultimately show "(\<Sum>\<^isub>\<infinity>n. measure P (F n)) = measure P (\<Union>i. F i)"
-      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_psuminf[symmetric]
+      using F by auto
+    ultimately show "(\<Sum>n. measure P (F n)) = measure P (\<Union>i. F i)"
+      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric]
                     M2.measure_countably_additive
                cong: M1.positive_integral_cong)
   qed
 
   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
-  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<omega>)"
+  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<infinity>)"
   proof (rule exI[of _ F], intro conjI)
     show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
     show "(\<Union>i. F i) = space P"
-      using F by (auto simp: pair_measure_def pair_measure_generator_def isoton_def)
-    show "\<forall>i. measure P (F i) \<noteq> \<omega>" using F by auto
+      using F by (auto simp: pair_measure_def pair_measure_generator_def)
+    show "\<forall>i. measure P (F i) \<noteq> \<infinity>" using F by auto
   qed
 qed
 
@@ -741,7 +744,7 @@
     show "(\<lambda>(y, x). (x, y)) \<in> measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable)
     show "sets (sigma E) = sets P" "space E = space P" "A \<in> sets (sigma E)"
       using assms unfolding pair_measure_def by auto
-    show "range F \<subseteq> sets E" "F \<up> space E" "\<And>i. \<mu> (F i) \<noteq> \<omega>"
+    show "range F \<subseteq> sets E" "incseq F" "(\<Union>i. F i) = space E" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
       using F `A \<in> sets P` by (auto simp: pair_measure_def)
     fix X assume "X \<in> sets E"
     then obtain A B where X[simp]: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
@@ -758,8 +761,8 @@
 qed
 
 lemma pair_sigma_algebra_sigma:
-  assumes 1: "S1 \<up> (space E1)" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
-  assumes 2: "S2 \<up> (space E2)" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
+  assumes 1: "incseq S1" "(\<Union>i. S1 i) = space E1" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
+  assumes 2: "decseq S2" "(\<Union>i. S2 i) = space E2" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
   shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
     (is "sets ?S = sets ?E")
 proof -
@@ -771,22 +774,22 @@
     using E1 E2 by (intro sigma_algebra_sigma) auto
   { fix A assume "A \<in> sets E1"
     then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
-      using E1 2 unfolding isoton_def pair_measure_generator_def by auto
+      using E1 2 unfolding pair_measure_generator_def by auto
     also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
     also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
       using 2 `A \<in> sets E1`
       by (intro sigma_sets.Union)
-         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+         (force simp: image_subset_iff intro!: sigma_sets.Basic)
     finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
   moreover
   { fix B assume "B \<in> sets E2"
     then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
-      using E2 1 unfolding isoton_def pair_measure_generator_def by auto
+      using E2 1 unfolding pair_measure_generator_def by auto
     also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
     also have "\<dots> \<in> sets ?E"
       using 1 `B \<in> sets E2` unfolding pair_measure_generator_def sets_sigma
       by (intro sigma_sets.Union)
-         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+         (force simp: image_subset_iff intro!: sigma_sets.Basic)
     finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
   ultimately have proj:
     "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
@@ -819,12 +822,12 @@
 section "Fubinis theorem"
 
 lemma (in pair_sigma_finite) simple_function_cut:
-  assumes f: "simple_function P f"
+  assumes f: "simple_function P f" "\<And>x. 0 \<le> f x"
   shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
 proof -
   have f_borel: "f \<in> borel_measurable P"
-    using f by (rule borel_measurable_simple_function)
+    using f(1) by (rule borel_measurable_simple_function)
   let "?F z" = "f -` {z} \<inter> space P"
   let "?F' x z" = "Pair x -` ?F z"
   { fix x assume "x \<in> space M1"
@@ -836,15 +839,15 @@
       by (intro borel_measurable_vimage measurable_cut_fst)
     ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
       apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
-      apply (rule simple_function_indicator_representation[OF f])
+      apply (rule simple_function_indicator_representation[OF f(1)])
       using `x \<in> space M1` by (auto simp del: space_sigma) }
   note M2_sf = this
   { fix x assume x: "x \<in> space M1"
     then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
-      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x]]
+      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
       unfolding simple_integral_def
     proof (safe intro!: setsum_mono_zero_cong_left)
-      from f show "finite (f ` space P)" by (rule simple_functionD)
+      from f(1) show "finite (f ` space P)" by (rule simple_functionD)
     next
       fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
         using `x \<in> space M1` by (auto simp: space_pair_measure)
@@ -862,11 +865,16 @@
     by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
   moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
     by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
+  moreover have *: "\<And>i x. 0 \<le> M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
+    using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst)
+  moreover { fix i assume "i \<in> f`space P"
+    with * have "\<And>x. 0 \<le> i * M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
+      using f(2) by auto }
   ultimately
   show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
     by (auto simp del: vimage_Int cong: measurable_cong
-             intro!: M1.borel_measurable_pextreal_setsum
+             intro!: M1.borel_measurable_extreal_setsum setsum_cong
              simp add: M1.positive_integral_setsum simple_integral_def
                        M1.positive_integral_cmult
                        M1.positive_integral_cong[OF eq]
@@ -880,42 +888,38 @@
       (is "?C f \<in> borel_measurable M1")
     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
 proof -
-  from borel_measurable_implies_simple_function_sequence[OF f]
-  obtain F where F: "\<And>i. simple_function P (F i)" "F \<up> f" by auto
+  from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
   then have F_borel: "\<And>i. F i \<in> borel_measurable P"
-    and F_mono: "\<And>i x. F i x \<le> F (Suc i) x"
-    and F_SUPR: "\<And>x. (SUP i. F i x) = f x"
-    unfolding isoton_fun_expand unfolding isoton_def le_fun_def
     by (auto intro: borel_measurable_simple_function)
-  note sf = simple_function_cut[OF F(1)]
+  note sf = simple_function_cut[OF F(1,5)]
   then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
     using F(1) by auto
   moreover
   { fix x assume "x \<in> space M1"
-    have isotone: "(\<lambda> i y. F i (x, y)) \<up> (\<lambda>y. f (x, y))"
-      using `F \<up> f` unfolding isoton_fun_expand
-      by (auto simp: isoton_def)
-    note measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
-    from M2.positive_integral_isoton[OF isotone this]
-    have "(SUP i. ?C (F i) x) = ?C f x"
-      by (simp add: isoton_def) }
+    from F measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
+    have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
+      by (intro M2.positive_integral_monotone_convergence_SUP)
+         (auto simp: incseq_Suc_iff le_fun_def)
+    then have "(SUP i. ?C (F i) x) = ?C f x"
+      unfolding F(4) positive_integral_max_0 by simp }
   note SUPR_C = this
   ultimately show "?C f \<in> borel_measurable M1"
     by (simp cong: measurable_cong)
   have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
-    using F_borel F_mono
-    by (auto intro!: positive_integral_monotone_convergence_SUP[symmetric])
+    using F_borel F
+    by (intro positive_integral_monotone_convergence_SUP) auto
   also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
     unfolding sf(2) by simp
-  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1"
-    by (auto intro!: M1.positive_integral_monotone_convergence_SUP[OF _ sf(1)]
-                     M2.positive_integral_mono F_mono)
+  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
+    by (intro M1.positive_integral_monotone_convergence_SUP[symmetric])
+       (auto intro!: M2.positive_integral_mono M2.positive_integral_positive
+                simp: incseq_Suc_iff le_fun_def)
   also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
-    using F_borel F_mono
-    by (auto intro!: M2.positive_integral_monotone_convergence_SUP
-                     M1.positive_integral_cong measurable_pair_image_snd)
+    using F_borel F(2,5)
+    by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric]
+             simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd)
   finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
-    unfolding F_SUPR by simp
+    using F by (simp add: positive_integral_max_0)
 qed
 
 lemma (in pair_sigma_finite) measure_preserving_swap:
@@ -963,8 +967,8 @@
   unfolding positive_integral_fst_measurable[OF assms] ..
 
 lemma (in pair_sigma_finite) AE_pair:
-  assumes "almost_everywhere (\<lambda>x. Q x)"
-  shows "M1.almost_everywhere (\<lambda>x. M2.almost_everywhere (\<lambda>y. Q (x, y)))"
+  assumes "AE x in P. Q x"
+  shows "AE x in M1. (AE y in M2. Q (x, y))"
 proof -
   obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
     using assms unfolding almost_everywhere_def by auto
@@ -972,9 +976,9 @@
   proof (rule M1.AE_I)
     from N measure_cut_measurable_fst[OF `N \<in> sets P`]
     show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
-      by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def)
+      by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
     show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro M1.borel_measurable_pextreal_neq_const measure_cut_measurable_fst N)
+      by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N)
     { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
       have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
       proof (rule M2.AE_I)
@@ -1036,41 +1040,52 @@
   shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
     and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
 proof -
-  let "?pf x" = "Real (f x)" and "?nf x" = "Real (- f x)"
+  let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)"
   have
     borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
-    int: "integral\<^isup>P P ?nf \<noteq> \<omega>" "integral\<^isup>P P ?pf \<noteq> \<omega>"
+    int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
     using assms by auto
-  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<omega>"
-     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<omega>"
+  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
     using borel[THEN positive_integral_fst_measurable(1)] int
     unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
   with borel[THEN positive_integral_fst_measurable(1)]
-  have AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<noteq> \<omega>)"
-    "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. Real (- f (x, y)) \<partial>M2) \<noteq> \<omega>)"
-    by (auto intro!: M1.positive_integral_omega_AE)
-  then show ?AE using assms
+  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+    "AE x in M1. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+    by (auto intro!: M1.positive_integral_PInf_AE )
+  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+    "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+    by (auto simp: M2.positive_integral_positive)
+  from AE_pos show ?AE using assms
     by (simp add: measurable_pair_image_snd integrable_def)
-  { fix f assume borel: "(\<lambda>x. Real (f x)) \<in> borel_measurable P"
-      and int: "integral\<^isup>P P (\<lambda>x. Real (f x)) \<noteq> \<omega>"
-      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2) \<noteq> \<omega>)"
-    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. Real (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
+  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+      using M2.positive_integral_positive
+      by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder)
+    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
+  note this[simp]
+  { fix f assume borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable P"
+      and int: "integral\<^isup>P P (\<lambda>x. extreal (f x)) \<noteq> \<infinity>"
+      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
+    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
     proof (intro integrable_def[THEN iffD2] conjI)
       show "?f \<in> borel_measurable M1"
-        using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable)
-      have "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. Real (f (x, y))  \<partial>M2) \<partial>M1)"
-        using AE by (auto intro!: M1.positive_integral_cong_AE simp: Real_real)
-      then show "(\<integral>\<^isup>+x. Real (?f x) \<partial>M1) \<noteq> \<omega>"
+        using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable)
+      have "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y))  \<partial>M2) \<partial>M1)"
+        using AE M2.positive_integral_positive
+        by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real)
+      then show "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) \<noteq> \<infinity>"
         using positive_integral_fst_measurable[OF borel] int by simp
-      have "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
-        by (intro M1.positive_integral_cong) simp
-      then show "(\<integral>\<^isup>+x. Real (- ?f x) \<partial>M1) \<noteq> \<omega>" by simp
+      have "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+        by (intro M1.positive_integral_cong_pos)
+           (simp add: M2.positive_integral_positive real_of_extreal_pos)
+      then show "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
     qed }
-  from this[OF borel(1) int(1) AE(2)] this[OF borel(2) int(2) AE(1)]
+  with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
   show ?INT
     unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
       borel[THEN positive_integral_fst_measurable(2), symmetric]
-    using AE by (simp add: M1.integral_real)
+    using AE[THEN M1.integral_real]
+    by simp
 qed
 
 lemma (in pair_sigma_finite) integrable_snd_measurable:
@@ -1195,7 +1210,8 @@
 
 lemma sigma_product_algebra_sigma_eq:
   assumes "finite I"
-  assumes isotone: "\<And>i. i \<in> I \<Longrightarrow> (S i) \<up> (space (E i))"
+  assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
+  assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (E i)"
   assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)"
   and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))"
   shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)"
@@ -1214,13 +1230,13 @@
     by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem)
   { fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
     then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E"
-      using isotone unfolding isoton_def space_product_algebra
+      using mono union unfolding incseq_Suc_iff space_product_algebra
       by (auto dest: Pi_mem)
     also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))"
       unfolding space_product_algebra
       apply simp
       apply (subst Pi_UN[OF `finite I`])
-      using isotone[THEN isoton_mono_le] apply simp
+      using mono[THEN incseqD] apply simp
       apply (simp add: PiE_Int)
       apply (intro PiE_cong)
       using A sets_into by (auto intro!: into_space)
@@ -1324,7 +1340,7 @@
     obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" .
     then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
       using sets_into_space `I \<inter> J = {}`
-      by (auto simp add: space_pair_measure) blast+
+      by (auto simp add: space_pair_measure) fast+
     then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)"
       using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI
         simp: product_algebra_def)
@@ -1360,6 +1376,44 @@
 sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra
   by default (fact finite_index')
 
+lemma setprod_extreal_0:
+  fixes f :: "'a \<Rightarrow> extreal"
+  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
+proof cases
+  assume "finite A"
+  then show ?thesis by (induct A) auto
+qed auto
+
+lemma setprod_extreal_pos:
+  fixes f :: "'a \<Rightarrow> extreal" assumes pos: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" shows "0 \<le> (\<Prod>i\<in>I. f i)"
+proof cases
+  assume "finite I" from this pos show ?thesis by induct auto
+qed simp
+
+lemma setprod_PInf:
+  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
+  shows "(\<Prod>i\<in>I. f i) = \<infinity> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
+proof cases
+  assume "finite I" from this assms show ?thesis
+  proof (induct I)
+    case (insert i I)
+    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_extreal_pos)
+    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
+    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
+      using setprod_extreal_pos[of I f] pos
+      by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto
+    also have "\<dots> \<longleftrightarrow> finite (insert i I) \<and> (\<exists>j\<in>insert i I. f j = \<infinity>) \<and> (\<forall>j\<in>insert i I. f j \<noteq> 0)"
+      using insert by (auto simp: setprod_extreal_0)
+    finally show ?case .
+  qed simp
+qed simp
+
+lemma setprod_extreal: "(\<Prod>i\<in>A. extreal (f i)) = extreal (setprod f A)"
+proof cases
+  assume "finite A" then show ?thesis
+    by induct (auto simp: one_extreal_def)
+qed (simp add: one_extreal_def)
+
 lemma (in finite_product_sigma_finite) product_algebra_generator_measure:
   assumes "Pi\<^isub>E I F \<in> sets G"
   shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))"
@@ -1373,13 +1427,13 @@
 next
   assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
   then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
-    by (auto simp: setprod_pextreal_0 intro!: bexI)
+    by (auto simp: setprod_extreal_0 intro!: bexI)
   moreover
   have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
     by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
        (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
   then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
-    by (auto simp: setprod_pextreal_0 intro!: bexI)
+    by (auto simp: setprod_extreal_0 intro!: bexI)
   ultimately show ?thesis
     unfolding product_algebra_generator_def by simp
 qed
@@ -1387,34 +1441,32 @@
 lemma (in finite_product_sigma_finite) sigma_finite_pairs:
   "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
     (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
-    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<omega>) \<and>
-    (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<up> space G"
+    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
+    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G"
 proof -
-  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> F \<up> space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<omega>)"
+  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)"
     using M.sigma_finite_up by simp
   from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-  then have "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. F i \<up> space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<omega>"
+  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>"
     by auto
   let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
   note space_product_algebra[simp]
   show ?thesis
-  proof (intro exI[of _ F] conjI allI isotoneI set_eqI iffI ballI)
+  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
     fix i show "range (F i) \<subseteq> sets (M i)" by fact
   next
-    fix i k show "\<mu> i (F i k) \<noteq> \<omega>" by fact
+    fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact
   next
     fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G"
       using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space
       by (force simp: image_subset_iff)
   next
     fix f assume "f \<in> space G"
-    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"]
-      `\<And>i. F i \<up> space (M i)`[THEN isotonD(2)]
-      `\<And>i. F i \<up> space (M i)`[THEN isoton_mono_le]
-    show "f \<in> (\<Union>i. ?F i)" by auto
+    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
+    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def)
   next
     fix i show "?F i \<subseteq> ?F (Suc i)"
-      using `\<And>i. F i \<up> space (M i)`[THEN isotonD(1)] by auto
+      using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto
   qed
 qed
 
@@ -1438,7 +1490,7 @@
 using `finite I` proof induct
   case empty
   interpret finite_product_sigma_finite M "{}" by default simp
-  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> pextreal"
+  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> extreal"
   show ?case
   proof (intro exI conjI ballI)
     have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
@@ -1488,22 +1540,23 @@
     proof (unfold *, default, simp)
       from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
       then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
-        "(\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) \<up> space I'.G"
-        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<omega>"
+        "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
+        "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G"
+        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>"
         by blast+
       let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k"
       show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
-          (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<omega>)"
+          (\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)"
       proof (intro exI[of _ ?F] conjI allI)
         show "range ?F \<subseteq> sets I'.P" using F(1) by auto
       next
-        from F(2)[THEN isotonD(2)]
-        show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
+        from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
       next
         fix j
-        show "?\<nu> (?F j) \<noteq> \<omega>"
-          using F `finite I`
-          by (subst product) (auto simp: setprod_\<omega>)
+        have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)"
+          using F(1) by auto
+        with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>"
+          by (subst product) auto
       qed
     qed
   qed
@@ -1542,7 +1595,8 @@
   by (simp_all add: sigma_def image_constant)
 
 lemma (in product_sigma_finite) positive_integral_empty:
-  "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
+  assumes pos: "0 \<le> f (\<lambda>k. undefined)"
+  shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
 proof -
   interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
   have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
@@ -1550,10 +1604,10 @@
   then show ?thesis
     unfolding positive_integral_def simple_function_def simple_integral_def_raw
   proof (simp add: P_empty, intro antisym)
-    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> f}. f (\<lambda>k. undefined))"
-      by (intro le_SUPI) auto
-    show "(SUP f:{g. g \<le> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)"
-      by (intro SUP_leI) (auto simp: le_fun_def)
+    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
+      by (intro le_SUPI) (auto simp: le_fun_def split: split_max)
+    show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
+      by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm)
   qed
 qed
 
@@ -1572,8 +1626,9 @@
   let "?X S" = "?g -` S \<inter> space P.P"
   from IJ.sigma_finite_pairs obtain F where
     F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
-       "(\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) \<up> space IJ.G"
-       "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<omega>"
+       "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
+       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G"
+       "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>"
     by auto
   let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
   show "IJ.\<mu> A = P.\<mu> (?X A)"
@@ -1589,12 +1644,13 @@
     show "range ?F \<subseteq> sets IJ.G" using F
       by (simp add: image_subset_iff product_algebra_def
                     product_algebra_generator_def)
-    show "?F \<up> space IJ.G " using F(2) by simp
-    show "\<And>k. IJ.\<mu> (?F k) \<noteq> \<omega>"
-      using `finite I` F
-      by (subst IJ.measure_times) (auto simp add: setprod_\<omega>)
-    show "?g \<in> measurable P.P IJ.P"
-      using IJ by (rule measurable_merge)
+    show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+
+  next
+    fix k
+    have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)"
+      using F(1) by auto
+    with F `finite I` setprod_PInf[of "I \<union> J", OF this]
+    show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
   next
     fix A assume "A \<in> sets IJ.G"
     then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F"
@@ -1611,7 +1667,7 @@
       using `finite J` `finite I` F unfolding A
       by (intro IJ.measure_times[symmetric]) auto
     finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym)
-  qed
+  qed (rule measurable_merge[OF IJ])
 qed
 
 lemma (in product_sigma_finite) measure_preserving_merge:
@@ -1692,8 +1748,9 @@
 qed
 
 lemma (in product_sigma_finite) product_positive_integral_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> pextreal"
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> extreal"
   assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
+  and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
 using assms proof induct
   case empty
@@ -1707,12 +1764,15 @@
     using insert by (auto intro!: setprod_cong)
   have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
     using sets_into_space insert
-    by (intro sigma_algebra.borel_measurable_pextreal_setprod sigma_algebra_product_algebra
+    by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra
               measurable_comp[OF measurable_component_singleton, unfolded comp_def])
        auto
-  show ?case
-    by (simp add: product_positive_integral_insert[OF insert(1,2) prod])
-       (simp add: insert I.positive_integral_cmult M.positive_integral_multc * prod subset_insertI)
+  then show ?case
+    apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
+    apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc)
+    apply (subst I.positive_integral_cmult)
+    apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive)
+    done
 qed
 
 lemma (in product_sigma_finite) product_integral_singleton:
@@ -1720,8 +1780,8 @@
   shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
 proof -
   interpret I: finite_product_sigma_finite M "{i}" by default simp
-  have *: "(\<lambda>x. Real (f x)) \<in> borel_measurable (M i)"
-    "(\<lambda>x. Real (- f x)) \<in> borel_measurable (M i)"
+  have *: "(\<lambda>x. extreal (f x)) \<in> borel_measurable (M i)"
+    "(\<lambda>x. extreal (- f x)) \<in> borel_measurable (M i)"
     using assms by auto
   show ?thesis
     unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
@@ -1795,15 +1855,17 @@
   proof (unfold integrable_def, intro conjI)
     show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
       using borel by auto
-    have "(\<integral>\<^isup>+x. Real (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. Real (abs (f i (x i)))) \<partial>P)"
-      by (simp add: Real_setprod abs_setprod)
-    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. Real (abs (f i x)) \<partial>M i))"
+    have "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. extreal (abs (f i (x i)))) \<partial>P)"
+      by (simp add: setprod_extreal abs_setprod)
+    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. extreal (abs (f i x)) \<partial>M i))"
       using f by (subst product_positive_integral_setprod) auto
-    also have "\<dots> < \<omega>"
+    also have "\<dots> < \<infinity>"
       using integrable[THEN M.integrable_abs]
-      unfolding pextreal_less_\<omega> setprod_\<omega> integrable_def by simp
-    finally show "(\<integral>\<^isup>+x. Real (abs (?f x)) \<partial>P) \<noteq> \<omega>" by auto
-    show "(\<integral>\<^isup>+x. Real (- abs (?f x)) \<partial>P) \<noteq> \<omega>" by simp
+      by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
+    finally show "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
+    have "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
+      by (intro positive_integral_cong_pos) auto
+    then show "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
   qed
   ultimately show ?thesis
     by (rule integrable_abs_iff[THEN iffD1])
--- a/src/HOL/Probability/Radon_Nikodym.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -2,87 +2,81 @@
 imports Lebesgue_Integration
 begin
 
-lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
-proof safe
-  assume "x < \<omega>"
-  then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
-  moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
-  ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
-qed auto
-
 lemma (in sigma_finite_measure) Ex_finite_integrable_function:
-  shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
+  shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
 proof -
   obtain A :: "nat \<Rightarrow> 'a set" where
     range: "range A \<subseteq> sets M" and
     space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
+    measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" and
     disjoint: "disjoint_family A"
     using disjoint_sigma_finite by auto
   let "?B i" = "2^Suc i * \<mu> (A i)"
   have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
   proof
-    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
-    proof cases
-      assume "\<mu> (A i) = 0"
-      then show ?thesis by (auto intro!: exI[of _ 1])
-    next
-      assume not_0: "\<mu> (A i) \<noteq> 0"
-      then have "?B i \<noteq> \<omega>" using measure[of i] by auto
-      then have "inverse (?B i) \<noteq> 0" unfolding pextreal_inverse_eq_0 by simp
-      then show ?thesis using measure[of i] not_0
-        by (auto intro!: exI[of _ "inverse (?B i) / 2"]
-                 simp: pextreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
-    qed
+    fix i have Ai: "A i \<in> sets M" using range by auto
+    from measure positive_measure[OF this]
+    show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
+      by (auto intro!: extreal_dense simp: extreal_0_gt_inverse)
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
     "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
-  let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
+  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
+  let "?h x" = "\<Sum>i. n i * indicator (A i) x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?h] del: notI)
     have "\<And>i. A i \<in> sets M"
       using range by fastsimp+
-    then have "integral\<^isup>P M ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
-      by (simp add: positive_integral_psuminf positive_integral_cmult_indicator)
-    also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
-    proof (rule psuminf_le)
-      fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
+    then have "integral\<^isup>P M ?h = (\<Sum>i. n i * \<mu> (A i))" using pos
+      by (simp add: positive_integral_suminf positive_integral_cmult_indicator)
+    also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)"
+    proof (rule suminf_le_pos)
+      fix N
+      have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)"
+        using positive_measure[OF `A N \<in> sets M`] n[of N]
+        by (intro extreal_mult_right_mono) auto
+      also have "\<dots> \<le> (1 / 2) ^ Suc N"
         using measure[of N] n[of N]
-        by (cases "n N")
-           (auto simp: pextreal_noteq_omega_Ex field_simps zero_le_mult_iff
-                       mult_le_0_iff mult_less_0_iff power_less_zero_eq
-                       power_le_zero_eq inverse_eq_divide less_divide_eq
-                       power_divide split: split_if_asm)
+        by (cases rule: extreal2_cases[of "n N" "\<mu> (A N)"])
+           (simp_all add: inverse_eq_divide power_divide one_extreal_def extreal_power_divide)
+      finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" .
+      show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp
     qed
-    also have "\<dots> = Real 1"
-      by (rule suminf_imp_psuminf, rule power_half_series, auto)
-    finally show "integral\<^isup>P M ?h \<noteq> \<omega>" by auto
+    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_extreal by auto
   next
-    fix x assume "x \<in> space M"
-    then obtain i where "x \<in> A i" using space[symmetric] by auto
-    from psuminf_cmult_indicator[OF disjoint, OF this]
-    have "?h x = n i" by simp
-    then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
+    { fix x assume "x \<in> space M"
+      then obtain i where "x \<in> A i" using space[symmetric] by auto
+      with disjoint n have "?h x = n i"
+        by (auto intro!: suminf_cmult_indicator intro: less_imp_le)
+      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by auto }
+    note pos = this
+    fix x show "0 \<le> ?h x"
+    proof cases
+      assume "x \<in> space M" then show "0 \<le> ?h x" using pos by (auto intro: less_imp_le)
+    next
+      assume "x \<notin> space M" then have "\<And>i. x \<notin> A i" using space by auto
+      then show "0 \<le> ?h x" by auto
+    qed
   next
-    show "?h \<in> borel_measurable M" using range
-      by (auto intro!: borel_measurable_psuminf borel_measurable_pextreal_times)
+    show "?h \<in> borel_measurable M" using range n
+      by (auto intro!: borel_measurable_psuminf borel_measurable_extreal_times extreal_0_le_mult intro: less_imp_le)
   qed
 qed
 
 subsection "Absolutely continuous"
 
 definition (in measure_space)
-  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pextreal))"
+  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: extreal))"
 
 lemma (in measure_space) absolutely_continuous_AE:
   assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
     and "absolutely_continuous (measure M')" "AE x. P x"
-  shows "measure_space.almost_everywhere M' P"
+   shows "AE x in M'. P x"
 proof -
   interpret \<nu>: measure_space M' by fact
   from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N"
     unfolding almost_everywhere_def by auto
-  show "\<nu>.almost_everywhere P"
+  show "AE x in M'. P x"
   proof (rule \<nu>.AE_I')
     show "{x\<in>space M'. \<not> P x} \<subseteq> N" by simp fact
     from `absolutely_continuous (measure M')` show "N \<in> \<nu>.null_sets"
@@ -99,7 +93,7 @@
   interpret v: finite_measure_space ?\<nu> by fact
   have "\<nu> N = measure ?\<nu> (\<Union>x\<in>N. {x})" by simp
   also have "\<dots> = (\<Sum>x\<in>N. measure ?\<nu> {x})"
-  proof (rule v.measure_finitely_additive''[symmetric])
+  proof (rule v.measure_setsum[symmetric])
     show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
     show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
     fix x assume "x \<in> N" thus "{x} \<in> sets ?\<nu>" using `N \<subseteq> space M` sets_eq_Pow by auto
@@ -107,8 +101,10 @@
   also have "\<dots> = 0"
   proof (safe intro!: setsum_0')
     fix x assume "x \<in> N"
-    hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
-    hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
+    hence "\<mu> {x} \<le> \<mu> N" "0 \<le> \<mu> {x}"
+      using sets_eq_Pow `N \<subseteq> space M` positive_measure[of "{x}"]
+      by (auto intro!: measure_mono)
+    then have "\<mu> {x} = 0" using `\<mu> N = 0` by simp
     thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
   qed
   finally show "\<nu> N = 0" by simp
@@ -125,12 +121,12 @@
 lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
   fixes e :: real assumes "0 < e"
   assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)"
-  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<le>
-                    real (\<mu> A) - real (\<nu> A) \<and>
-                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (\<nu> B))"
+  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
+                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
+                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
 proof -
-  let "?d A" = "real (\<mu> A) - real (\<nu> A)"
   interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
+  let "?d A" = "\<mu>' A - M'.\<mu>' A"
   let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
     then {}
     else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
@@ -157,7 +153,7 @@
       fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
       hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
       hence "?d (A n \<union> B) = ?d (A n) + ?d B"
-        using `A n \<in> sets M` real_finite_measure_Union M'.real_finite_measure_Union by simp
+        using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by simp
       also have "\<dots> \<le> ?d (A n) - e" using dB by simp
       finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
     qed }
@@ -186,11 +182,7 @@
         fix n assume "?d (space M) \<le> ?d (space M - A n)"
         also have "\<dots> \<le> ?d (space M - A (Suc n))"
           using A_in_sets sets_into_space dA_mono[of n]
-            real_finite_measure_Diff[of "space M"]
-            real_finite_measure_Diff[of "space M"]
-            M'.real_finite_measure_Diff[of "space M"]
-            M'.real_finite_measure_Diff[of "space M"]
-          by (simp del: A_simps)
+          by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff)
         finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
       qed simp
     qed
@@ -200,13 +192,16 @@
     { fix n have "?d (A n) \<le> - real n * e"
       proof (induct n)
         case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
-      qed simp } note dA_less = this
+      next
+        case 0 with M'.empty_measure show ?case by (simp add: zero_extreal_def)
+      qed } note dA_less = this
     have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
     proof (rule incseq_SucI)
       fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto
     qed
-    from real_finite_continuity_from_below[of A] `range A \<subseteq> sets M`
-      M'.real_finite_continuity_from_below[of A]
+    have A: "incseq A" by (auto intro!: incseq_SucI)
+    from finite_continuity_from_below[OF _ A] `range A \<subseteq> sets M`
+      M'.finite_continuity_from_below[OF _ A]
     have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
       by (auto intro!: LIMSEQ_diff)
     obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto
@@ -216,33 +211,55 @@
   qed
 qed
 
+lemma (in finite_measure) restricted_measure_subset:
+  assumes S: "S \<in> sets M" and X: "X \<subseteq> S"
+  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
+proof cases
+  note r = restricted_finite_measure[OF S]
+  { assume "X \<in> sets M" with S X show ?thesis
+      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def by auto }
+  { assume "X \<notin> sets M"
+    moreover then have "S \<inter> X \<notin> sets M"
+      using X by (simp add: Int_absorb1)
+    ultimately show ?thesis
+      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def using S by auto }
+qed
+
+lemma (in finite_measure) restricted_measure:
+  assumes X: "S \<in> sets M" "X \<in> sets (restricted_space S)"
+  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
+proof -
+  from X have "S \<in> sets M" "X \<subseteq> S" by auto
+  from restricted_measure_subset[OF this] show ?thesis .
+qed
+
 lemma (in finite_measure) Radon_Nikodym_aux:
   assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
-  shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<le>
-                    real (\<mu> A) - real (\<nu> A) \<and>
-                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (\<nu> B))"
+  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
+                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
+                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
 proof -
-  let "?d A" = "real (\<mu> A) - real (\<nu> A)"
-  let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
   interpret M': finite_measure ?M' where
     "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto
+  let "?d A" = "\<mu>' A - M'.\<mu>' A"
+  let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
   let "?r S" = "restricted_space S"
-  { fix S n
-    assume S: "S \<in> sets M"
-    hence **: "\<And>X. X \<in> op \<inter> S ` sets M \<longleftrightarrow> X \<in> sets M \<and> X \<subseteq> S" by auto
-    have [simp]: "(restricted_space S\<lparr>measure := \<nu>\<rparr>) = M'.restricted_space S"
-      by (cases M) simp
-    from M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S
-    have "finite_measure (?r S)" "0 < 1 / real (Suc n)"
+  { fix S n assume S: "S \<in> sets M"
+    note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S
+    then have "finite_measure (?r S)" "0 < 1 / real (Suc n)"
       "finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto
-    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X ..
-    hence "?P X S n"
-    proof (simp add: **, safe)
-      fix C assume C: "C \<in> sets M" "C \<subseteq> X" "X \<subseteq> S" and
-        *: "\<forall>B\<in>sets M. S \<inter> B \<subseteq> X \<longrightarrow> - (1 / real (Suc n)) < ?d (S \<inter> B)"
-      hence "C \<subseteq> S" "C \<subseteq> X" "S \<inter> C = C" by auto
-      with *[THEN bspec, OF `C \<in> sets M`]
-      show "- (1 / real (Suc n)) < ?d C" by auto
+    from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this
+    have "?P X S n"
+    proof (intro conjI ballI impI)
+      show "X \<in> sets M" "X \<subseteq> S" using X(1) `S \<in> sets M` by auto
+      have "S \<in> op \<inter> S ` sets M" using `S \<in> sets M` by auto
+      then show "?d S \<le> ?d X"
+        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp
+      fix C assume "C \<in> sets M" "C \<subseteq> X"
+      then have "C \<in> sets (restricted_space S)" "C \<subseteq> X"
+        using `S \<in> sets M` `X \<subseteq> S` by auto
+      with X(2) show "- 1 / real (Suc n) < ?d C"
+        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto
     qed
     hence "\<exists>A. ?P A S n" by auto }
   note Ex_P = this
@@ -268,10 +285,11 @@
   show ?thesis
   proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
     show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
-    from `range A \<subseteq> sets M` A_mono
-      real_finite_continuity_from_above[of A]
-      M'.real_finite_continuity_from_above[of A]
-    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: LIMSEQ_diff)
+    have A: "decseq A" using A_mono by (auto intro!: decseq_SucI)
+    from
+      finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
+      M'.finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
+    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (intro LIMSEQ_diff)
     thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
       by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
   next
@@ -290,6 +308,10 @@
   qed
 qed
 
+lemma (in finite_measure) real_measure:
+  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = extreal r"
+  using finite_measure[OF A] positive_measure[OF A] by (cases "\<mu> A") auto
+
 lemma (in finite_measure) Radon_Nikodym_finite_measure:
   assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
   assumes "absolutely_continuous \<nu>"
@@ -298,7 +320,7 @@
   interpret M': finite_measure ?M'
     where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>"
     using assms(1) by auto
-  def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A}"
+  def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A)}"
   have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
   hence "G \<noteq> {}" by auto
   { fix f g assume f: "f \<in> G" and g: "g \<in> G"
@@ -324,24 +346,28 @@
       also have "\<dots> = \<nu> A"
         using M'.measure_additive[OF sets] union by auto
       finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" .
+    next
+      fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max)
     qed }
   note max_in_G = this
-  { fix f g assume  "f \<up> g" and f: "\<And>i. f i \<in> G"
-    have "g \<in> G" unfolding G_def
+  { fix f assume  "incseq f" and f: "\<And>i. f i \<in> G"
+    have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def
     proof safe
-      from `f \<up> g` have [simp]: "g = (\<lambda>x. SUP i. f i x)"
-        unfolding isoton_def fun_eq_iff SUPR_apply by simp
-      have f_borel: "\<And>i. f i \<in> borel_measurable M" using f unfolding G_def by simp
-      thus "g \<in> borel_measurable M" by auto
+      show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M"
+        using f by (auto simp: G_def)
+      { fix x show "0 \<le> (SUP i. f i x)"
+          using f by (auto simp: G_def intro: le_SUPI2) }
+    next
       fix A assume "A \<in> sets M"
-      hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M"
-        using f_borel by (auto intro!: borel_measurable_indicator)
-      from positive_integral_isoton[OF isoton_indicator[OF `f \<up> g`] this]
-      have SUP: "(\<integral>\<^isup>+x. g x * indicator A x \<partial>M) =
-          (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))"
-        unfolding isoton_def by simp
-      show "(\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A" unfolding SUP
-        using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI)
+      have "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
+        (\<integral>\<^isup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
+        by (intro positive_integral_cong) (simp split: split_indicator)
+      also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))"
+        using `incseq f` f `A \<in> sets M`
+        by (intro positive_integral_monotone_convergence_SUP)
+           (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
+      finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> \<nu> A"
+        using f `A \<in> sets M` by (auto intro!: SUP_leI simp: G_def)
     qed }
   note SUP_in_G = this
   let ?y = "SUP g : G. integral\<^isup>P M g"
@@ -351,9 +377,8 @@
     from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)"
       by (simp cong: positive_integral_cong)
   qed
-  hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto
-  from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = this
-  hence "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n"
+  from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^isup>P M"] guess ys .. note ys = this
+  then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n"
   proof safe
     fix n assume "range ys \<subseteq> integral\<^isup>P M ` G"
     hence "ys n \<in> integral\<^isup>P M ` G" by auto
@@ -362,8 +387,9 @@
   from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^isup>P M (gs n) = ys n" by auto
   hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto
   let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
-  def f \<equiv> "SUP i. ?g i"
-  have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
+  def f \<equiv> "\<lambda>x. SUP i. ?g i x"
+  let "?F A x" = "f x * indicator A x"
+  have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
   { fix i have "?g i \<in> G"
     proof (induct i)
       case 0 thus ?case by simp fact
@@ -373,15 +399,13 @@
         by (auto simp add: atMost_Suc intro!: max_in_G)
     qed }
   note g_in_G = this
-  have "\<And>x. \<forall>i. ?g i x \<le> ?g (Suc i) x"
-    using gs_not_empty by (simp add: atMost_Suc)
-  hence isoton_g: "?g \<up> f" by (simp add: isoton_def le_fun_def f_def)
-  from SUP_in_G[OF this g_in_G] have "f \<in> G" .
-  hence [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
-  have "(\<lambda>i. integral\<^isup>P M (?g i)) \<up> integral\<^isup>P M f"
-    using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def)
-  hence "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))"
-    unfolding isoton_def by simp
+  have "incseq ?g" using gs_not_empty
+    by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)
+  from SUP_in_G[OF this g_in_G] have "f \<in> G" unfolding f_def .
+  then have [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
+  have "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding f_def
+    using g_in_G `incseq ?g`
+    by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
   also have "\<dots> = ?y"
   proof (rule antisym)
     show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
@@ -390,42 +414,57 @@
       by (auto intro!: SUP_mono positive_integral_mono Max_ge)
   qed
   finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
-  let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+  have "\<And>x. 0 \<le> f x"
+    unfolding f_def using `\<And>i. gs i \<in> G`
+    by (auto intro!: le_SUPI2 Max_ge_iff[THEN iffD2] simp: G_def)
+  let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
   let ?M = "M\<lparr> measure := ?t\<rparr>"
   interpret M: sigma_algebra ?M
     by (intro sigma_algebra_cong) auto
+  have f_le_\<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> \<nu> A"
+    using `f \<in> G` unfolding G_def by auto
   have fmM: "finite_measure ?M"
-  proof (default, simp_all add: countably_additive_def, safe)
+  proof (default, simp_all add: countably_additive_def positive_def, safe del: notI)
     fix A :: "nat \<Rightarrow> 'a set"  assume A: "range A \<subseteq> sets M" "disjoint_family A"
-    have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x \<partial>M))
-      = (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x) \<partial>M)"
-      using `range A \<subseteq> sets M`
-      by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
-    also have "\<dots> = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x \<partial>M)"
-      apply (rule positive_integral_cong)
-      apply (subst psuminf_cmult_right)
-      unfolding psuminf_indicator[OF `disjoint_family A`] ..
-    finally have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) x \<partial>M))
-      = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) x \<partial>M)" .
-    moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
+    have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. (\<Sum>n. ?F (A n) x) \<partial>M)"
+      using `range A \<subseteq> sets M` `\<And>x. 0 \<le> f x`
+      by (intro positive_integral_suminf[symmetric]) auto
+    also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)"
+      using `\<And>x. 0 \<le> f x`
+      by (intro positive_integral_cong) (simp add: suminf_cmult_extreal suminf_indicator[OF `disjoint_family A`])
+    finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" .
+    moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
       using M'.measure_countably_additive A by (simp add: comp_def)
-    moreover have "\<And>i. (\<integral>\<^isup>+x. f x * indicator (A i) x \<partial>M) \<le> \<nu> (A i)"
-        using A `f \<in> G` unfolding G_def by auto
-    moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN)
+    moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<infinity>" using M'.finite_measure A by (simp add: countable_UN)
     moreover {
-      have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)"
+      have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)"
         using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
-      also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>)
-      finally have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<noteq> \<omega>"
-        by (simp add: pextreal_less_\<omega>) }
+      also have "\<nu> (\<Union>i. A i) < \<infinity>" using v_fin by simp
+      finally have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<noteq> \<infinity>" by simp }
+    moreover have "\<And>i. (\<integral>\<^isup>+x. ?F (A i) x \<partial>M) \<le> \<nu> (A i)"
+      using A by (intro f_le_\<nu>) auto
     ultimately
-    show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
-      apply (subst psuminf_minus) by simp_all
+    show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)"
+      by (subst suminf_extreal_minus) (simp_all add: positive_integral_positive)
+  next
+    fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M"
+      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def extreal_le_minus_iff)
+  next
+    show "\<nu> (space M) - (\<integral>\<^isup>+ x. ?F (space M) x \<partial>M) \<noteq> \<infinity>" (is "?a - ?b \<noteq> _")
+      using positive_integral_positive[of "?F (space M)"]
+      by (cases rule: extreal2_cases[of ?a ?b]) auto
   qed
   then interpret M: finite_measure ?M
     where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
     by (simp_all add: fmM)
-  have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto
+  have ac: "absolutely_continuous ?t" unfolding absolutely_continuous_def
+  proof
+    fix N assume N: "N \<in> null_sets"
+    with `absolutely_continuous \<nu>` have "\<nu> N = 0" unfolding absolutely_continuous_def by auto
+    moreover with N have "(\<integral>\<^isup>+ x. ?F N x \<partial>M) \<le> \<nu> N" using `f \<in> G` by (auto simp: G_def)
+    ultimately show "\<nu> N - (\<integral>\<^isup>+ x. ?F N x \<partial>M) = 0"
+      using positive_integral_positive by (auto intro!: antisym)
+  qed
   have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
   proof (rule ccontr)
     assume "\<not> ?thesis"
@@ -436,43 +475,54 @@
       using M.measure_mono[of A "space M"] A sets_into_space by simp
     finally have pos_t: "0 < ?t (space M)" by simp
     moreover
-    hence pos_M: "0 < \<mu> (space M)"
-      using ac top unfolding absolutely_continuous_def by auto
+    then have "\<mu> (space M) \<noteq> 0"
+      using ac unfolding absolutely_continuous_def by auto
+    then have pos_M: "0 < \<mu> (space M)"
+      using positive_measure[OF top] by (simp add: le_less)
     moreover
     have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)"
       using `f \<in> G` unfolding G_def by auto
-    hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<omega>"
+    hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
       using M'.finite_measure_of_space by auto
     moreover
     def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
-    ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>"
-      using M'.finite_measure_of_space
-      by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space)
+    ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
+      using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"]
+      by (cases rule: extreal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
+         (simp_all add: field_simps)
+    then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
     let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
     interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto
-    have "finite_measure ?Mb"
-      by default
-         (insert finite_measure_of_space b measure_countably_additive,
-          auto simp: psuminf_cmult_right countably_additive_def)
+    have Mb: "finite_measure ?Mb"
+    proof
+      show "positive ?Mb (measure ?Mb)"
+        using `0 \<le> b` by (auto simp: positive_def)
+      show "countably_additive ?Mb (measure ?Mb)"
+        using `0 \<le> b` measure_countably_additive
+        by (auto simp: countably_additive_def suminf_cmult_extreal subset_eq)
+      show "measure ?Mb (space ?Mb) \<noteq> \<infinity>"
+        using b by auto
+    qed
     from M.Radon_Nikodym_aux[OF this]
     obtain A0 where "A0 \<in> sets M" and
       space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
-      *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)" by auto
-    { fix B assume "B \<in> sets M" "B \<subseteq> A0"
+      *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)"
+      unfolding M.\<mu>'_def finite_measure.\<mu>'_def[OF Mb] by auto
+    { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
       with *[OF this] have "b * \<mu> B \<le> ?t B"
-        using M'.finite_measure b finite_measure
-        by (cases "b * \<mu> B", cases "?t B") (auto simp: field_simps) }
+        using M'.finite_measure b finite_measure M.positive_measure[OF B(1)]
+        by (cases rule: extreal2_cases[of "?t B" "b * \<mu> B"]) auto }
     note bM_le_t = this
     let "?f0 x" = "f x + b * indicator A0 x"
     { fix A assume A: "A \<in> sets M"
       hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
       have "(\<integral>\<^isup>+x. ?f0 x  * indicator A x \<partial>M) =
         (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
-        by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith)
+        by (auto intro!: positive_integral_cong split: split_indicator)
       hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
           (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)"
-        using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A
-        by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) }
+        using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`
+        by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) }
     note f0_eq = this
     { fix A assume A: "A \<in> sets M"
       hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
@@ -487,39 +537,57 @@
         using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
         by (auto intro!: add_left_mono)
       also have "\<dots> \<le> \<nu> A"
-        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`]
-        by (cases "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M)", cases "\<nu> A", auto)
+        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`] positive_integral_positive[of "?F A"]
+        by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") auto
       finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
-    hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
-      by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times)
-    have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>"
-      "b * \<mu> (space M) \<noteq> \<omega>" "b * \<mu> A0 \<noteq> \<omega>"
+    hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
+      by (auto intro!: borel_measurable_indicator borel_measurable_extreal_add
+                       borel_measurable_extreal_times extreal_add_nonneg_nonneg)
+    have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>"
+      "b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>"
       using `A0 \<in> sets M` b
         finite_measure[of A0] M.finite_measure[of A0]
         finite_measure_of_space M.finite_measure_of_space
       by auto
-    have int_f_finite: "integral\<^isup>P M f \<noteq> \<omega>"
-      using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff
+    have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
+      using M'.finite_measure_of_space pos_t unfolding extreal_less_minus_iff
       by (auto cong: positive_integral_cong)
-    have "?t (space M) > b * \<mu> (space M)" unfolding b_def
-      apply (simp add: field_simps)
-      apply (subst mult_assoc[symmetric])
-      apply (subst pextreal_mult_inverse)
+    have  "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def
       using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
-      using pextreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"]
-      by simp_all
-    hence  "0 < ?t (space M) - b * \<mu> (space M)"
-      by (simp add: pextreal_zero_less_diff_iff)
+      using positive_integral_positive[of "?F (space M)"]
+      by (cases rule: extreal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
+         (auto simp: field_simps mult_less_cancel_left)
     also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
-      using space_less_A0 pos_M pos_t b real[unfolded pextreal_noteq_omega_Ex] by auto
-    finally have "b * \<mu> A0 < ?t A0" unfolding pextreal_zero_less_diff_iff .
-    hence "0 < ?t A0" by auto
-    hence "0 < \<mu> A0" using ac unfolding absolutely_continuous_def
+      using space_less_A0 b
+      using
+        `A0 \<in> sets M`[THEN M.real_measure]
+        top[THEN M.real_measure]
+      apply safe
+      apply simp
+      using
+        `A0 \<in> sets M`[THEN real_measure]
+        `A0 \<in> sets M`[THEN M'.real_measure]
+        top[THEN real_measure]
+        top[THEN M'.real_measure]
+      by (cases b) auto
+    finally have 1: "b * \<mu> A0 < ?t A0"
+      using
+        `A0 \<in> sets M`[THEN M.real_measure]
+      apply safe
+      apply simp
+      using
+        `A0 \<in> sets M`[THEN real_measure]
+        `A0 \<in> sets M`[THEN M'.real_measure]
+      by (cases b) auto
+    have "0 < ?t A0"
+      using b `A0 \<in> sets M` by (auto intro!: le_less_trans[OF _ 1])
+    then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def
       using `A0 \<in> sets M` by auto
-    hence "0 < b * \<mu> A0" using b by auto
-    from int_f_finite this
-    have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
-      by (rule pextreal_less_add)
+    then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto
+    hence "0 < b * \<mu> A0" using b by (auto simp: extreal_zero_less_0_iff)
+    with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
+      using `f \<in> G`
+      by (intro extreal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive)
     also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
       by (simp cong: positive_integral_cong)
     finally have "?y < integral\<^isup>P M ?f0" by simp
@@ -528,14 +596,15 @@
   qed
   show ?thesis
   proof (safe intro!: bexI[of _ f])
-    fix A assume "A\<in>sets M"
+    fix A assume A: "A\<in>sets M"
     show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
     proof (rule antisym)
       show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
         using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
       show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
         using upper_bound[THEN bspec, OF `A \<in> sets M`]
-         by (simp add: pextreal_zero_le_diff)
+        using M'.real_measure[OF A]
+        by (cases "integral\<^isup>P M (?F A)") auto
     qed
   qed simp
 qed
@@ -543,22 +612,22 @@
 lemma (in finite_measure) split_space_into_finite_sets_and_rest:
   assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
   assumes ac: "absolutely_continuous \<nu>"
-  shows "\<exists>\<Omega>0\<in>sets M. \<exists>\<Omega>::nat\<Rightarrow>'a set. disjoint_family \<Omega> \<and> range \<Omega> \<subseteq> sets M \<and> \<Omega>0 = space M - (\<Union>i. \<Omega> i) \<and>
-    (\<forall>A\<in>sets M. A \<subseteq> \<Omega>0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<omega>)) \<and>
-    (\<forall>i. \<nu> (\<Omega> i) \<noteq> \<omega>)"
+  shows "\<exists>A0\<in>sets M. \<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> A0 = space M - (\<Union>i. B i) \<and>
+    (\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<infinity>)) \<and>
+    (\<forall>i. \<nu> (B i) \<noteq> \<infinity>)"
 proof -
   interpret v: measure_space ?N
     where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
     by fact auto
-  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
+  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<infinity>}"
   let ?a = "SUP Q:?Q. \<mu> Q"
   have "{} \<in> ?Q" using v.empty_measure by auto
   then have Q_not_empty: "?Q \<noteq> {}" by blast
   have "?a \<le> \<mu> (space M)" using sets_into_space
     by (auto intro!: SUP_leI measure_mono top)
-  then have "?a \<noteq> \<omega>" using finite_measure_of_space
+  then have "?a \<noteq> \<infinity>" using finite_measure_of_space
     by auto
-  from SUPR_countable_SUPR[OF this Q_not_empty]
+  from SUPR_countable_SUPR[OF Q_not_empty, of \<mu>]
   obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
     by auto
   then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto
@@ -569,7 +638,7 @@
   have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
   proof (rule continuity_from_below[of ?O])
     show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
-    show "\<And>i. ?O i \<subseteq> ?O (Suc i)" by fastsimp
+    show "incseq ?O" by (fastsimp intro!: incseq_SucI)
   qed
   have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
   have O_sets: "\<And>i. ?O i \<in> sets M"
@@ -580,8 +649,8 @@
       using Q' by (auto intro: finite_UN)
     with v.measure_finitely_subadditive[of "{.. i}" Q']
     have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
-    also have "\<dots> < \<omega>" unfolding setsum_\<omega> pextreal_less_\<omega> using Q' by auto
-    finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pextreal_less_\<omega> by auto
+    also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)
+    finally show "\<nu> (?O i) \<noteq> \<infinity>" by simp
   qed auto
   have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp
   have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
@@ -592,7 +661,7 @@
     proof (safe intro!: Sup_mono, unfold bex_simps)
       fix i
       have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
-      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<omega>) \<and>
+      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<infinity>) \<and>
         \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
         using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
     qed
@@ -610,50 +679,52 @@
     show "range Q \<subseteq> sets M"
       using Q_sets by auto
     { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
-      show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
+      show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>"
       proof (rule disjCI, simp)
-        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<omega>"
+        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<infinity>"
         show "\<mu> A = 0 \<and> \<nu> A = 0"
         proof cases
           assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0"
             unfolding absolutely_continuous_def by auto
           ultimately show ?thesis by simp
         next
-          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<omega>" by auto
+          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<infinity>" using positive_measure[OF A(1)] by auto
           with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
             using Q' by (auto intro!: measure_additive countable_UN)
           also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
           proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
             show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
-              using `\<nu> A \<noteq> \<omega>` O_sets A by auto
-          qed fastsimp
+              using `\<nu> A \<noteq> \<infinity>` O_sets A by auto
+          qed (fastsimp intro!: incseq_SucI)
           also have "\<dots> \<le> ?a"
-          proof (safe intro!: SUPR_bound)
+          proof (safe intro!: SUP_leI)
             fix i have "?O i \<union> A \<in> ?Q"
             proof (safe del: notI)
               show "?O i \<union> A \<in> sets M" using O_sets A by auto
               from O_in_G[of i]
               moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
                 using v.measure_subadditive[of "?O i" A] A O_sets by auto
-              ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
-                using `\<nu> A \<noteq> \<omega>` by auto
+              ultimately show "\<nu> (?O i \<union> A) \<noteq> \<infinity>"
+                using `\<nu> A \<noteq> \<infinity>` by auto
             qed
             then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
           qed
-          finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
-            by (cases "\<mu> A") (auto simp: pextreal_noteq_omega_Ex)
+          finally have "\<mu> A = 0"
+            unfolding a_eq using real_measure[OF `?O_0 \<in> sets M`] real_measure[OF A(1)] by auto
           with `\<mu> A \<noteq> 0` show ?thesis by auto
         qed
       qed }
-    { fix i show "\<nu> (Q i) \<noteq> \<omega>"
+    { fix i show "\<nu> (Q i) \<noteq> \<infinity>"
       proof (cases i)
         case 0 then show ?thesis
           unfolding Q_def using Q'[of 0] by simp
       next
         case (Suc n)
         then show ?thesis unfolding Q_def
-          using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
-          using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
+          using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
+          using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"]
+          using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono]
+          by (cases rule: extreal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
       qed }
     show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
     { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
@@ -675,7 +746,7 @@
 lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
   assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
   assumes "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
 proof -
   interpret v: measure_space ?N
     where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
@@ -684,14 +755,14 @@
   obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
     where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
     and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
-    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
-    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<omega>" by force
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>"
+    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<infinity>" by force
   from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
-  have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
+  have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
     \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
   proof
     fix i
-    have indicator_eq: "\<And>f x A. (f x :: pextreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
+    have indicator_eq: "\<And>f x A. (f x :: extreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
       = (f x * indicator (Q i) x) * indicator A x"
       unfolding indicator_def by auto
     have fm: "finite_measure (restricted_space (Q i))"
@@ -702,7 +773,7 @@
     proof
       show "measure_space ?Q"
         using v.restricted_measure_space Q_sets[of i] by auto
-      show "measure ?Q (space ?Q) \<noteq> \<omega>" using Q_fin by simp
+      show "measure ?Q (space ?Q) \<noteq> \<infinity>" using Q_fin by simp
     qed
     have "R.absolutely_continuous \<nu>"
       using `absolutely_continuous \<nu>` `Q i \<in> sets M`
@@ -712,48 +783,40 @@
       and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \<partial>M)"
       unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
         positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
-    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
+    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
       \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
-      by (fastsimp intro!: exI[of _ "\<lambda>x. f x * indicator (Q i) x"] positive_integral_cong
-          simp: indicator_def)
+      by (auto intro!: exI[of _ "\<lambda>x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos
+        split: split_indicator split_if_asm simp: max_def)
   qed
-  from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M"
+  from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
     and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
       \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)"
     by auto
-  let "?f x" =
-    "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator Q0 x"
+  let "?f x" = "(\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?f])
-    show "?f \<in> borel_measurable M"
-      by (safe intro!: borel_measurable_psuminf borel_measurable_pextreal_times
-        borel_measurable_pextreal_add borel_measurable_indicator
-        borel_measurable_const borel Q_sets Q0 Diff countable_UN)
+    show "?f \<in> borel_measurable M" using Q0 borel Q_sets
+      by (auto intro!: measurable_If)
+    show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def)
     fix A assume "A \<in> sets M"
-    have *:
-      "\<And>x i. indicator A x * (f i x * indicator (Q i) x) =
-        f i x * indicator (Q i \<inter> A) x"
-      "\<And>x i. (indicator A x * indicator Q0 x :: pextreal) =
-        indicator (Q0 \<inter> A) x" by (auto simp: indicator_def)
-    have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) =
-      (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)"
-      unfolding f[OF `A \<in> sets M`]
-      apply (simp del: pextreal_times(2) add: field_simps *)
-      apply (subst positive_integral_add)
-      apply (fastsimp intro: Q0 `A \<in> sets M`)
-      apply (fastsimp intro: Q_sets `A \<in> sets M` borel_measurable_psuminf borel)
-      apply (subst positive_integral_cmult_indicator)
-      apply (fastsimp intro: Q0 `A \<in> sets M`)
-      unfolding psuminf_cmult_right[symmetric]
-      apply (subst positive_integral_psuminf)
-      apply (fastsimp intro: `A \<in> sets M` Q_sets borel)
-      apply (simp add: *)
-      done
-    moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
+    have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
+    have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
+      "\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x"
+      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_extreal_times)
+    have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
+      using borel by (intro positive_integral_cong) (auto simp: indicator_def)
+    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)"
+      using borel Qi Q0(1) `A \<in> sets M`
+      by (subst positive_integral_add) (auto simp del: extreal_infty_mult
+          simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
+    also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)"
+      by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto
+    finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)" .
+    moreover have "(\<Sum>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
       using Q Q_sets `A \<in> sets M`
       by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
          (auto simp: disjoint_family_on_def)
-    moreover have "\<omega> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)"
+    moreover have "\<infinity> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)"
     proof -
       have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
       from in_Q0[OF this] show ?thesis by auto
@@ -770,40 +833,43 @@
 
 lemma (in sigma_finite_measure) Radon_Nikodym:
   assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-  assumes "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+  assumes ac: "absolutely_continuous \<nu>"
+  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
 proof -
   from Ex_finite_integrable_function
-  obtain h where finite: "integral\<^isup>P M h \<noteq> \<omega>" and
+  obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and
     borel: "h \<in> borel_measurable M" and
+    nn: "\<And>x. 0 \<le> h x" and
     pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
-    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" by auto
+    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
   let "?T A" = "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
   let ?MT = "M\<lparr> measure := ?T \<rparr>"
-  from measure_space_density[OF borel] finite
   interpret T: finite_measure ?MT
     where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T"
-    unfolding finite_measure_def finite_measure_axioms_def
-    by (simp_all cong: positive_integral_cong)
-  have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pextreal)} = N"
-    using sets_into_space pos by (force simp: indicator_def)
-  then have "T.absolutely_continuous \<nu>" using assms(2) borel
-    unfolding T.absolutely_continuous_def absolutely_continuous_def
-    by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff)
+    unfolding finite_measure_def finite_measure_axioms_def using borel finite nn
+    by (auto intro!: measure_space_density cong: positive_integral_cong)
+  have "T.absolutely_continuous \<nu>"
+  proof (unfold T.absolutely_continuous_def, safe)
+    fix N assume "N \<in> sets M" "(\<integral>\<^isup>+x. h x * indicator N x \<partial>M) = 0"
+    with borel ac pos have "AE x. x \<notin> N"
+      by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le)
+    then have "N \<in> null_sets" using `N \<in> sets M` sets_into_space
+      by (subst (asm) AE_iff_measurable[OF `N \<in> sets M`]) auto
+    then show "\<nu> N = 0" using ac by (auto simp: absolutely_continuous_def)
+  qed
   from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
-  obtain f where f_borel: "f \<in> borel_measurable M" and
+  obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and
     fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>?MT)"
     by (auto simp: measurable_def)
   show ?thesis
   proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
     show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
-      using borel f_borel by (auto intro: borel_measurable_pextreal_times)
+      using borel f_borel by (auto intro: borel_measurable_extreal_times)
+    show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto
     fix A assume "A \<in> sets M"
-    then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
-      using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator)
-    from positive_integral_translated_density[OF borel this]
-    show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
-      unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
+    then show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
+      unfolding fT[OF `A \<in> sets M`] mult_assoc using nn borel f_borel
+      by (intro positive_integral_translated_density) auto
   qed
 qed
 
@@ -811,7 +877,8 @@
 
 lemma (in measure_space) finite_density_unique:
   assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  and fin: "integral\<^isup>P M f < \<omega>"
+  assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
+  and fin: "integral\<^isup>P M f \<noteq> \<infinity>"
   shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M))
     \<longleftrightarrow> (AE x. f x = g x)"
     (is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _")
@@ -822,42 +889,38 @@
 next
   assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
   from this[THEN bspec, OF top] fin
-  have g_fin: "integral\<^isup>P M g < \<omega>" by (simp cong: positive_integral_cong)
+  have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
   { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-      and g_fin: "integral\<^isup>P M g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+      and pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
+      and g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
     let ?N = "{x\<in>space M. g x < f x}"
     have N: "?N \<in> sets M" using borel by simp
+    have "?P g ?N \<le> integral\<^isup>P M g" using pos
+      by (intro positive_integral_mono_AE) (auto split: split_indicator)
+    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by auto
     have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
       by (auto intro!: positive_integral_cong simp: indicator_def)
     also have "\<dots> = ?P f ?N - ?P g ?N"
     proof (rule positive_integral_diff)
       show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
         using borel N by auto
-      have "?P g ?N \<le> integral\<^isup>P M g"
-        by (auto intro!: positive_integral_mono simp: indicator_def)
-      then show "?P g ?N \<noteq> \<omega>" using g_fin by auto
-      fix x assume "x \<in> space M"
-      show "g x * indicator ?N x \<le> f x * indicator ?N x"
-        by (auto simp: indicator_def)
-    qed
+      show "AE x. g x * indicator ?N x \<le> f x * indicator ?N x"
+           "AE x. 0 \<le> g x * indicator ?N x"
+        using pos by (auto split: split_indicator)
+    qed fact
     also have "\<dots> = 0"
-      using eq[THEN bspec, OF N] by simp
-    finally have "\<mu> {x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = 0"
-      using borel N by (subst (asm) positive_integral_0_iff) auto
-    moreover have "{x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = ?N"
-      by (auto simp: pextreal_zero_le_diff)
-    ultimately have "?N \<in> null_sets" using N by simp }
-  from this[OF borel g_fin eq] this[OF borel(2,1) fin]
-  have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} \<in> null_sets"
-    using eq by (intro null_sets_Un) auto
-  also have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} = {x\<in>space M. f x \<noteq> g x}"
-    by auto
-  finally show "AE x. f x = g x"
-    unfolding almost_everywhere_def by auto
+      unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto
+    finally have "AE x. f x \<le> g x"
+      using pos borel positive_integral_PInf_AE[OF borel(2) g_fin]
+      by (subst (asm) positive_integral_0_iff_AE)
+         (auto split: split_indicator simp: not_less extreal_minus_le_iff) }
+  from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
+  show "AE x. f x = g x" by auto
 qed
 
 lemma (in finite_measure) density_unique_finite_measure:
   assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
+  assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> f' x"
   assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
     (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
   shows "AE x. f x = f' x"
@@ -865,26 +928,26 @@
   let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A"
   let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x"
   interpret M: measure_space "M\<lparr> measure := ?\<nu>\<rparr>"
-    using borel(1) by (rule measure_space_density) simp
+    using borel(1) pos(1) by (rule measure_space_density) simp
   have ac: "absolutely_continuous ?\<nu>"
     using f by (rule density_is_absolutely_continuous)
   from split_space_into_finite_sets_and_rest[OF `measure_space (M\<lparr> measure := ?\<nu>\<rparr>)` ac]
   obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
     where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
     and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
-    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<omega>"
-    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<omega>" by force
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<infinity>"
+    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<infinity>" by force
   from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
   let ?N = "{x\<in>space M. f x \<noteq> f' x}"
   have "?N \<in> sets M" using borel by auto
-  have *: "\<And>i x A. \<And>y::pextreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+  have *: "\<And>i x A. \<And>y::extreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
     unfolding indicator_def by auto
-  have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q
+  have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
     by (intro finite_density_unique[THEN iffD1] allI)
-       (auto intro!: borel_measurable_pextreal_times f Int simp: *)
+       (auto intro!: borel_measurable_extreal_times f Int simp: *)
   moreover have "AE x. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
-    { fix f :: "'a \<Rightarrow> pextreal" assume borel: "f \<in> borel_measurable M"
+    { fix f :: "'a \<Rightarrow> extreal" assume borel: "f \<in> borel_measurable M"
         and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
       let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
       have "(\<Union>i. ?A i) \<in> null_sets"
@@ -896,69 +959,74 @@
           by (auto intro!: positive_integral_mono simp: indicator_def)
         also have "\<dots> = of_nat i * \<mu> (?A i)"
           using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
-        also have "\<dots> < \<omega>"
+        also have "\<dots> < \<infinity>"
           using `?A i \<in> sets M`[THEN finite_measure] by auto
-        finally have "?\<nu> (?A i) \<noteq> \<omega>" by simp
+        finally have "?\<nu> (?A i) \<noteq> \<infinity>" by simp
         then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
       qed
-      also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x < \<omega>}"
-        by (auto simp: less_\<omega>_Ex_of_nat)
-      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" by (simp add: pextreal_less_\<omega>) }
+      also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
+        by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat)
+      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" by simp }
     from this[OF borel(1) refl] this[OF borel(2) f]
-    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>} \<in> null_sets" by simp_all
-    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>}) \<in> null_sets" by (rule null_sets_Un)
+    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets" by simp_all
+    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets" by (rule null_sets_Un)
     show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq>
-      (Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>})" by (auto simp: indicator_def)
+      (Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def)
   qed
   moreover have "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
     ?f (space M) x = ?f' (space M) x"
     by (auto simp: indicator_def Q0)
   ultimately have "AE x. ?f (space M) x = ?f' (space M) x"
-    by (auto simp: all_AE_countable)
+    by (auto simp: AE_all_countable[symmetric])
   then show "AE x. f x = f' x" by auto
 qed
 
 lemma (in sigma_finite_measure) density_unique:
-  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
-  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
+  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
+  assumes f': "f' \<in> borel_measurable M" "AE x. 0 \<le> f' x"
+  assumes eq: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
     (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
   shows "AE x. f x = f' x"
 proof -
   obtain h where h_borel: "h \<in> borel_measurable M"
-    and fin: "integral\<^isup>P M h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>"
+    and fin: "integral\<^isup>P M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
     using Ex_finite_integrable_function by auto
-  interpret h: measure_space "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
-    using h_borel by (rule measure_space_density) simp
+  then have h_nn: "AE x. 0 \<le> h x" by auto
+  let ?H = "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
+  have H: "measure_space ?H"
+    using h_borel h_nn by (rule measure_space_density) simp
+  then interpret h: measure_space ?H .
   interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
     by default (simp cong: positive_integral_cong add: fin)
   let ?fM = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)\<rparr>"
   interpret f: measure_space ?fM
-    using borel(1) by (rule measure_space_density) simp
+    using f by (rule measure_space_density) simp
   let ?f'M = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)\<rparr>"
   interpret f': measure_space ?f'M
-    using borel(2) by (rule measure_space_density) simp
+    using f' by (rule measure_space_density) simp
   { fix A assume "A \<in> sets M"
-    then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pextreal)} = A"
-      using pos sets_into_space by (force simp: indicator_def)
+    then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
+      using pos(1) sets_into_space by (force simp: indicator_def)
     then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets"
-      using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) }
+      using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto }
   note h_null_sets = this
   { fix A assume "A \<in> sets M"
-    have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
-      using `A \<in> sets M` h_borel borel
-      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
+    have "(\<integral>\<^isup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
+      using `A \<in> sets M` h_borel h_nn f f'
+      by (intro positive_integral_translated_density[symmetric]) auto
     also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
-      by (rule f'.positive_integral_cong_measure) (simp_all add: f)
-    also have "\<dots> = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)"
-      using `A \<in> sets M` h_borel borel
-      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
-    finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)" . }
-  then have "h.almost_everywhere (\<lambda>x. f x = f' x)"
-    using h_borel borel
-    apply (intro h.density_unique_finite_measure)
-    apply (simp add: measurable_def)
-    apply (simp add: measurable_def)
-    by (simp add: positive_integral_translated_density)
+      by (rule f'.positive_integral_cong_measure) (simp_all add: eq)
+    also have "\<dots> = (\<integral>\<^isup>+x. f' x * (h x * indicator A x) \<partial>M)"
+      using `A \<in> sets M` h_borel h_nn f f'
+      by (intro positive_integral_translated_density) auto
+    finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)"
+      by (simp add: ac_simps)
+    then have "(\<integral>\<^isup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^isup>+x. (f' x * indicator A x) \<partial>?H)"
+      using `A \<in> sets M` h_borel h_nn f f'
+      by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto }
+  then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
+    by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous)
+       simp_all
   then show "AE x. f x = f' x"
     unfolding h.almost_everywhere_def almost_everywhere_def
     by (auto simp add: h_null_sets)
@@ -966,41 +1034,42 @@
 
 lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
   assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-    and f: "f \<in> borel_measurable M"
+    and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
     and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-  shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)"
+  shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<infinity>)"
 proof
   assume "sigma_finite_measure ?N"
   then interpret \<nu>: sigma_finite_measure ?N
     where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
     and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def)
   from \<nu>.Ex_finite_integrable_function obtain h where
-    h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<omega>"
-    and fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>" by auto
-  have "AE x. f x * h x \<noteq> \<omega>"
+    h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<infinity>" and
+    h_nn: "\<And>x. 0 \<le> h x" and
+    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
+  have "AE x. f x * h x \<noteq> \<infinity>"
   proof (rule AE_I')
-    have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h
+    have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn
       by (subst \<nu>.positive_integral_cong_measure[symmetric,
           of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"])
          (auto intro!: positive_integral_translated_density simp: eq)
-    then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<omega>"
+    then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
       using h(2) by simp
-    then show "(\<lambda>x. f x * h x) -` {\<omega>} \<inter> space M \<in> null_sets"
-      using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage)
+    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets"
+      using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage)
   qed auto
-  then show "AE x. f x \<noteq> \<omega>"
+  then show "AE x. f x \<noteq> \<infinity>"
     using fin by (auto elim!: AE_Ball_mp)
 next
-  assume AE: "AE x. f x \<noteq> \<omega>"
+  assume AE: "AE x. f x \<noteq> \<infinity>"
   from sigma_finite guess Q .. note Q = this
   interpret \<nu>: measure_space ?N
     where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
     and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
-  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<omega>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
+  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
   { fix i j have "A i \<inter> Q j \<in> sets M"
     unfolding A_def using f Q
     apply (rule_tac Int)
-    by (cases i) (auto intro: measurable_sets[OF f]) }
+    by (cases i) (auto intro: measurable_sets[OF f(1)]) }
   note A_in_sets = this
   let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j"
   show "sigma_finite_measure ?N"
@@ -1021,18 +1090,21 @@
       fix x assume x: "x \<in> space M"
       show "x \<in> (\<Union>i. A i)"
       proof (cases "f x")
-        case infinite then show ?thesis using x unfolding A_def by (auto intro: exI[of _ 0])
+        case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
       next
-        case (preal r)
-        with less_\<omega>_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by auto
-        then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"])
+        case (real r)
+        with less_PInf_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by (auto simp: real_eq_of_nat)
+        then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"])
+      next
+        case MInf with x show ?thesis
+          unfolding A_def by (auto intro!: exI[of _ "Suc 0"])
       qed
     qed (auto simp: A_def)
     finally show "(\<Union>i. ?A i) = space ?N" by simp
   next
     fix n obtain i j where
       [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
-    have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<omega>"
+    have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
     proof (cases i)
       case 0
       have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
@@ -1045,11 +1117,11 @@
         by (auto intro!: positive_integral_mono simp: indicator_def A_def)
       also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
         using Q by (auto intro!: positive_integral_cmult_indicator)
-      also have "\<dots> < \<omega>"
-        using Q by auto
+      also have "\<dots> < \<infinity>"
+        using Q by (auto simp: real_eq_of_nat[symmetric])
       finally show ?thesis by simp
     qed
-    then show "measure ?N (?A n) \<noteq> \<omega>"
+    then show "measure ?N (?A n) \<noteq> \<infinity>"
       using A_in_sets Q eq by auto
   qed
 qed
@@ -1057,7 +1129,7 @@
 section "Radon-Nikodym derivative"
 
 definition
-  "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
+  "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and>
     (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
 
 lemma (in sigma_finite_measure) RN_deriv_cong:
@@ -1078,9 +1150,12 @@
   shows "RN_deriv M \<nu> \<in> borel_measurable M" (is ?borel)
   and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
     (is "\<And>A. _ \<Longrightarrow> ?int A")
+  and "0 \<le> RN_deriv M \<nu> x"
 proof -
   note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
-  thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
+  then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
+  from Ex show "0 \<le> RN_deriv M \<nu> x" unfolding RN_deriv_def
+    by (rule someI2_ex) simp
   fix A assume "A \<in> sets M"
   from Ex show "?int A" unfolding RN_deriv_def
     by (rule someI2_ex) (simp add: `A \<in> sets M`)
@@ -1092,22 +1167,28 @@
   shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
 proof -
   interpret \<nu>: measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f =
-    integral\<^isup>P (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>) f"
-    by (intro \<nu>.positive_integral_cong_measure[symmetric])
-       (simp_all add:  RN_deriv(2)[OF \<nu>, symmetric])
+  note RN = RN_deriv[OF \<nu>]
+  have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>)"
+    unfolding positive_integral_max_0 ..
+  also have "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>) =
+    (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>)"
+    by (intro \<nu>.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2))
+  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * max 0 (f x) \<partial>M)"
+    by (intro positive_integral_translated_density) (auto simp add: RN f)
   also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
-    by (intro positive_integral_translated_density)
-       (simp_all add: RN_deriv[OF \<nu>] f)
+    using RN_deriv(3)[OF \<nu>]
+    by (auto intro!: positive_integral_cong_pos split: split_if_asm
+             simp: max_def extreal_mult_le_0_iff)
   finally show ?thesis .
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_unique:
   assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
-  and f: "f \<in> borel_measurable M"
+  and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
   and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
   shows "AE x. f x = RN_deriv M \<nu> x"
 proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
+  show "AE x. 0 \<le> RN_deriv M \<nu> x" using RN_deriv[OF \<nu>] by auto
   fix A assume A: "A \<in> sets M"
   show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
     unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
@@ -1143,7 +1224,7 @@
   interpret M': sigma_finite_measure M'
   proof
     from sigma_finite guess F .. note F = this
-    show "\<exists>A::nat \<Rightarrow> 'c set. range A \<subseteq> sets M' \<and> (\<Union>i. A i) = space M' \<and> (\<forall>i. M'.\<mu> (A i) \<noteq> \<omega>)"
+    show "\<exists>A::nat \<Rightarrow> 'c set. range A \<subseteq> sets M' \<and> (\<Union>i. A i) = space M' \<and> (\<forall>i. M'.\<mu> (A i) \<noteq> \<infinity>)"
     proof (intro exI conjI allI)
       show *: "range (\<lambda>i. T' -` F i \<inter> space M') \<subseteq> sets M'"
         using F T' by (auto simp: measurable_def measure_preserving_def)
@@ -1157,7 +1238,7 @@
       then have "T -` (T' -` F i \<inter> space M') \<inter> space M = F i"
         using T inv sets_into_space[OF Fi]
         by (auto simp: measurable_def measure_preserving_def)
-      ultimately show "measure M' (T' -` F i \<inter> space M') \<noteq> \<omega>"
+      ultimately show "measure M' (T' -` F i \<inter> space M') \<noteq> \<infinity>"
         using F by simp
     qed
   qed
@@ -1165,6 +1246,7 @@
     by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+
   then show "(\<lambda>x. RN_deriv M' \<nu>' (T x)) \<in> borel_measurable M"
     by (simp add: comp_def)
+  show "AE x. 0 \<le> RN_deriv M' \<nu>' (T x)" using M'.RN_deriv(3)[OF \<nu>'] by auto
   fix A let ?A = "T' -` A \<inter> space M'"
   assume A: "A \<in> sets M"
   then have A': "?A \<in> sets M'" using T' unfolding measurable_def measure_preserving_def
@@ -1185,12 +1267,12 @@
 
 lemma (in sigma_finite_measure) RN_deriv_finite:
   assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<nu>"
-  shows "AE x. RN_deriv M \<nu> x \<noteq> \<omega>"
+  shows "AE x. RN_deriv M \<nu> x \<noteq> \<infinity>"
 proof -
   interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
   have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
   from sfm show ?thesis
-    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv[OF \<nu> ac]] by simp
+    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv(1)[OF \<nu> ac]] RN_deriv(2,3)[OF \<nu> ac] by simp
 qed
 
 lemma (in sigma_finite_measure)
@@ -1203,22 +1285,24 @@
 proof -
   interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
   have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-  have minus_cong: "\<And>A B A' B'::pextreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
+  have minus_cong: "\<And>A B A' B'::extreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
   have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
-  have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f unfolding measurable_def by auto
+  have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp
   { fix f :: "'a \<Rightarrow> real"
-    { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<omega>"
-      have "Real (real (RN_deriv M \<nu> x)) * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)"
+    { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>"
+      have "extreal (real (RN_deriv M \<nu> x)) * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
         by (simp add: mult_le_0_iff)
-      then have "RN_deriv M \<nu> x * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)"
-        using * by (simp add: Real_real) }
-    then have "(\<integral>\<^isup>+x. RN_deriv M \<nu> x * Real (f x) \<partial>M) = (\<integral>\<^isup>+x. Real (real (RN_deriv M \<nu> x) * f x) \<partial>M)"
-      using RN_deriv_finite[OF \<nu>] by (auto intro: positive_integral_cong_AE) }
-  with this this f f' Nf
+      then have "RN_deriv M \<nu> x * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
+        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: extreal_real split: split_if_asm) }
+    then have "(\<integral>\<^isup>+x. extreal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * extreal (f x) \<partial>M)"
+              "(\<integral>\<^isup>+x. extreal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * extreal (- f x) \<partial>M)"
+      using RN_deriv_finite[OF \<nu>] unfolding extreal_mult_minus_right uminus_extreal.simps(1)[symmetric]
+      by (auto intro!: positive_integral_cong_AE) }
+  note * = this
   show ?integral ?integrable
-    unfolding lebesgue_integral_def integrable_def
-    by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong
-             simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
+    unfolding lebesgue_integral_def integrable_def *
+    using f RN_deriv(1)[OF ms \<nu>(2)]
+    by (auto simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_singleton:
@@ -1231,7 +1315,7 @@
   from deriv(2)[OF `{x} \<in> sets M`]
   have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)"
     by (auto simp: indicator_def intro!: positive_integral_cong)
-  thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
+  thus ?thesis using positive_integral_cmult_indicator[OF _ `{x} \<in> sets M`] deriv(3)
     by auto
 qed
 
--- a/src/HOL/Probability/Sigma_Algebra.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,7 +1,7 @@
-(*  Title:      HOL/Probability/Sigma_Algebra.thy
-    Author:     Stefan Richter
-    Author:     Markus Wenzel
-    Author:     Lawrence Paulson
+(*  Title:      Sigma_Algebra.thy
+    Author:     Stefan Richter, Markus Wenzel, TU Muenchen
+    Plus material from the Hurd/Coble measure theory development,
+    translated by Lawrence Paulson.
 *)
 
 header {* Sigma Algebras *}
@@ -70,6 +70,16 @@
   "finite X \<Longrightarrow> X \<subseteq> sets M \<Longrightarrow> Union X \<in> sets M"
   by (induct set: finite) (auto simp add: Un)
 
+lemma (in algebra) finite_UN[intro]:
+  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
+  shows "(\<Union>i\<in>I. A i) \<in> sets M"
+  using assms by induct auto
+
+lemma (in algebra) finite_INT[intro]:
+  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
+  shows "(\<Inter>i\<in>I. A i) \<in> sets M"
+  using assms by (induct rule: finite_ne_induct) auto
+
 lemma algebra_iff_Int:
      "algebra M \<longleftrightarrow>
        sets M \<subseteq> Pow (space M) & {} \<in> sets M &
@@ -149,11 +159,6 @@
   ultimately show ?thesis by simp
 qed
 
-lemma (in sigma_algebra) finite_UN:
-  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
-  shows "(\<Union>i\<in>I. A i) \<in> sets M"
-  using assms by induct auto
-
 lemma (in sigma_algebra) countable_INT [intro]:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
   assumes A: "A`X \<subseteq> sets M" "X \<noteq> {}"
@@ -167,11 +172,6 @@
   ultimately show ?thesis by metis
 qed
 
-lemma (in sigma_algebra) finite_INT:
-  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
-  shows "(\<Inter>i\<in>I. A i) \<in> sets M"
-  using assms by (induct rule: finite_ne_induct) auto
-
 lemma algebra_Pow:
      "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
   by (auto simp add: algebra_def)
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -1,5 +1,5 @@
 theory Dining_Cryptographers
-imports Information
+imports "~~/src/HOL/Probability/Information"
 begin
 
 locale finite_space =
@@ -8,7 +8,7 @@
   and not_empty[simp]: "S \<noteq> {}"
 
 definition (in finite_space) "M = \<lparr> space = S, sets = Pow S,
-  measure = (\<lambda>A. of_nat (card A) / of_nat (card S) :: pextreal) \<rparr>"
+  measure = \<lambda>A. extreal (card A / card S) \<rparr>"
 
 lemma (in finite_space)
   shows space_M[simp]: "space M = S"
@@ -19,13 +19,14 @@
 proof (rule finite_measure_spaceI)
   fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M"
   then show "measure M (A \<union> B) = measure M A + measure M B"
-    by (simp add: inverse_eq_divide field_simps Real_real M_def
-                  divide_le_0_iff zero_le_divide_iff
-                  card_Un_disjoint finite_subset[OF _ finite])
-qed (auto simp: M_def)
+    by (simp add: M_def card_Un_disjoint finite_subset[OF _ finite] field_simps)
+qed (auto simp: M_def divide_nonneg_nonneg)
 
 sublocale finite_space \<subseteq> information_space M 2
-    by default (simp_all add: M_def)
+  by default (simp_all add: M_def one_extreal_def)
+
+lemma (in finite_space) \<mu>'_eq[simp]: "\<mu>' A = (if A \<subseteq> S then card A / card S else 0)"
+  unfolding \<mu>'_def by (auto simp: M_def)
 
 lemma set_of_list_extend:
   "{xs. length xs = Suc n \<and> (\<forall>x\<in>set xs. x \<in> A)} =
@@ -491,7 +492,7 @@
   let ?dI = "distribution inversion"
 
   { have "\<H>(inversion|payer) =
-        - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. real (?dIP {(x, z)}) * log 2 (real (?dIP {(x, z)}) / real (?dP {z}))))"
+        - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. ?dIP {(x, z)} * log 2 (?dIP {(x, z)} / ?dP {z})))"
       unfolding conditional_entropy_eq[OF simple_function_finite simple_function_finite]
       by (simp add: image_payer_dc_crypto setsum_Sigma)
     also have "... =
@@ -505,30 +506,27 @@
       moreover from x z obtain i where "z = Some i" and "i < n" by auto
       moreover from x have "length x = n" by (auto simp: inversion_def_raw dc_crypto)
       ultimately
-      have "real (?dIP {(x, z)}) = 2 / (real n * 2^n)" using x
-        unfolding distribution_def
-        by (simp add: M_def card_dc_crypto card_payer_and_inversion
-                      inverse_eq_divide mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
+      have "?dIP {(x, z)} = 2 / (real n * 2^n)" using x
+        by (auto simp add: card_dc_crypto card_payer_and_inversion distribution_def)
       moreover
       from z have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}"
         by (auto simp: dc_crypto payer_def)
       hence "card (payer -` {z} \<inter> dc_crypto) = 2^n"
         using card_list_length[where A="UNIV::bool set"]
         by (simp add: card_cartesian_product_singleton)
-      hence "real (?dP {z}) = 1 / real n" unfolding distribution_def
-        by (simp add: card_dc_crypto field_simps inverse_eq_divide M_def
-                      mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
+      hence "?dP {z} = 1 / real n"
+        by (simp add: distribution_def card_dc_crypto)
       ultimately
-      show "real (?dIP {(x,z)}) * log 2 (real (?dIP {(x,z)}) / real (?dP {z})) =
+      show "?dIP {(x,z)} * log 2 (?dIP {(x,z)} / ?dP {z}) =
        2 / (real n * 2^n) * (1 - real n)"
-        by (simp add: field_simps log_divide log_nat_power[of 2])
+        by (simp add: log_divide log_nat_power[of 2])
     qed
     also have "... = real n - 1"
       using n finite_space
       by (simp add: card_image_inversion card_image[OF inj_Some] field_simps real_eq_of_nat[symmetric])
     finally have "\<H>(inversion|payer) = real n - 1" . }
   moreover
-  { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. real (?dI {x}) * log 2 (real (?dI {x})))"
+  { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. ?dI {x} * log 2 (?dI {x}))"
       unfolding entropy_eq[OF simple_function_finite] by simp
     also have "... = - (\<Sum>x \<in> inversion`dc_crypto. 2 * (1 - real n) / 2^n)"
       unfolding neg_equal_iff_equal
@@ -536,10 +534,9 @@
       fix x assume x_inv: "x \<in> inversion ` dc_crypto"
       hence "length x = n" by (auto simp: inversion_def_raw dc_crypto)
       moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
-      ultimately have "?dI {x} = 2 / 2^n" using `0 < n` unfolding distribution_def
-        by (simp add: card_inversion[OF x_inv] card_dc_crypto M_def
-                      mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
-      thus "real (?dI {x}) * log 2 (real (?dI {x})) = 2 * (1 - real n) / 2^n"
+      ultimately have "?dI {x} = 2 / 2^n" using `0 < n`
+        by (auto simp: card_inversion[OF x_inv] card_dc_crypto distribution_def)
+      thus "?dI {x} * log 2 (?dI {x}) = 2 * (1 - real n) / 2^n"
         by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide)
     qed
     also have "... = real n - 1"
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -3,7 +3,7 @@
 header {* Formalization of a Countermeasure by Koepf \& Duermuth 2009 *}
 
 theory Koepf_Duermuth_Countermeasure
-  imports Information "~~/src/HOL/Library/Permutation"
+  imports "~~/src/HOL/Probability/Information" "~~/src/HOL/Library/Permutation"
 begin
 
 lemma
@@ -201,14 +201,17 @@
 lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
    by (auto intro!: setsum_nonneg)
 
-sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>"
+sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
   by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset)
 
-sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>"
+sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
+  by default (simp add: one_extreal_def)
+
+sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>" b
   by default simp
 
-sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. Real (setsum p S)\<rparr>" b
-  by default simp
+lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> \<mu>' A = setsum p A"
+  unfolding \<mu>'_def by auto
 
 locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b
     for b :: real
@@ -259,16 +262,6 @@
   "snd ` (SIGMA x:f`X. f -` {x} \<inter> X) = X"
   by (auto simp: image_iff)
 
-lemma zero_le_eq_True: "0 \<le> (x::pextreal) \<longleftrightarrow> True" by simp
-
-lemma Real_setprod:
-  assumes"\<And>i. i\<in>X \<Longrightarrow> 0 \<le> f i"
-  shows "(\<Prod>i\<in>X. Real (f i)) = Real (\<Prod>i\<in>X. f i)"
-proof cases
-  assume "finite X"
-  from this assms show ?thesis by induct (auto simp: mult_le_0_iff)
-qed simp
-
 lemma inj_Cons[simp]: "\<And>X. inj_on (\<lambda>(xs, x). x#xs) X" by (auto intro!: inj_onI)
 
 lemma setprod_setsum_distrib_lists:
@@ -323,10 +316,10 @@
  "p A \<equiv> setsum P A"
 
 abbreviation probability ("\<P>'(_') _") where
- "\<P>(X) x \<equiv> real (distribution X x)"
+ "\<P>(X) x \<equiv> distribution X x"
 
 abbreviation joint_probability ("\<P>'(_, _') _") where
- "\<P>(X, Y) x \<equiv> real (joint_distribution X Y x)"
+ "\<P>(X, Y) x \<equiv> joint_distribution X Y x"
 
 abbreviation conditional_probability ("\<P>'(_|_') _") where
  "\<P>(X|Y) x \<equiv> \<P>(X, Y) x / \<P>(Y) (snd`x)"
@@ -355,11 +348,12 @@
   from assms have *:
       "fst -` {k} \<inter> msgs = {k}\<times>{ms. length ms = n \<and> (\<forall>M\<in>set ms. M \<in> messages)}"
     unfolding msgs_def by auto
-  show "\<P>(fst) {k} = K k" unfolding distribution_def
-    apply (simp add: *) unfolding P_def
+  show "\<P>(fst) {k} = K k"
+    apply (simp add: \<mu>'_eq distribution_def)
+    apply (simp add: * P_def)
     apply (simp add: setsum_cartesian_product')
-    using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"]
-    by (simp add: setsum_right_distrib[symmetric] subset_eq setprod_1)
+    using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] `k \<in> keys`
+    by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod_1)
 qed
 
 lemma fst_image_msgs[simp]: "fst`msgs = keys"
@@ -390,7 +384,7 @@
 
       have "\<P>(OB, fst) {(obs, k)} / K k =
           p ({k}\<times>{ms. (k,ms) \<in> msgs \<and> OB (k,ms) = obs}) / K k"
-        unfolding distribution_def by (auto intro!: arg_cong[where f=p])
+        apply (simp add: distribution_def \<mu>'_eq) by (auto intro!: arg_cong[where f=p])
       also have "\<dots> =
           (\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
         unfolding P_def using `K k \<noteq> 0` `k \<in> keys`
@@ -415,7 +409,7 @@
       unfolding disjoint_family_on_def by auto
     have "\<P>(t\<circ>OB, fst) {(t obs, k)} = (\<Sum>obs'\<in>?S obs. \<P>(OB, fst) {(obs', k)})"
       unfolding distribution_def comp_def
-      using real_finite_measure_finite_Union[OF _ df]
+      using finite_measure_finite_Union[OF _ _ df]
       by (force simp add: * intro!: setsum_nonneg)
     also have "(\<Sum>obs'\<in>?S obs. \<P>(OB, fst) {(obs', k)}) = real (card (?S obs)) * \<P>(OB, fst) {(obs, k)}"
       by (simp add: t_eq_imp[OF `k \<in> keys` `K k \<noteq> 0` obs] real_eq_of_nat)
@@ -433,11 +427,11 @@
     then have "real (card ?S) \<noteq> 0" by auto
 
     have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / \<P>(t\<circ>OB) {t obs}"
-      using real_distribution_order'[of fst k "t\<circ>OB" "t obs"]
+      using distribution_order(7,8)[where X=fst and x=k and Y="t\<circ>OB" and y="t obs"]
       by (subst joint_distribution_commute) auto
     also have "\<P>(t\<circ>OB) {t obs} = (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'})"
-      using setsum_real_distribution(2)[of "t\<circ>OB" fst "t obs", symmetric]
-      using real_distribution_order'[of fst _ "t\<circ>OB" "t obs"] by (auto intro!: setsum_cong)
+      using setsum_distribution(2)[of "t\<circ>OB" fst "t obs", symmetric]
+      by (auto intro!: setsum_cong distribution_order(8))
     also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
       \<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
       using CP_t_K[OF `k\<in>keys` obs] CP_t_K[OF _ obs] `real (card ?S) \<noteq> 0`
@@ -445,11 +439,10 @@
           mult_divide_mult_cancel_left[OF `real (card ?S) \<noteq> 0`]
         cong: setsum_cong)
     also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
-      using setsum_real_distribution(2)[of OB fst obs, symmetric]
-      using real_distribution_order'[of fst _ OB obs] by (auto intro!: setsum_cong)
+      using setsum_distribution(2)[of OB fst obs, symmetric]
+      by (auto intro!: setsum_cong distribution_order(8))
     also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
-      using real_distribution_order'[of fst k OB obs]
-      by (subst joint_distribution_commute) auto
+      by (subst joint_distribution_commute) (auto intro!: distribution_order(8))
     finally have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}" . }
   note CP_T_eq_CP_O = this
 
@@ -472,7 +465,7 @@
       unfolding disjoint_family_on_def by auto
     have "\<P>(t\<circ>OB) {t (OB x)} = (\<Sum>obs\<in>?S (OB x). \<P>(OB) {obs})"
       unfolding distribution_def comp_def
-      using real_finite_measure_finite_Union[OF _ df]
+      using finite_measure_finite_Union[OF _ _ df]
       by (force simp add: * intro!: setsum_nonneg) }
   note P_t_sum_P_O = this
 
--- a/src/HOL/RealVector.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/RealVector.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -534,6 +534,34 @@
 lemma dist_triangle3: "dist x y \<le> dist a x + dist a y"
 using dist_triangle2 [of x y a] by (simp add: dist_commute)
 
+lemma dist_triangle_alt:
+  shows "dist y z <= dist x y + dist x z"
+by (rule dist_triangle3)
+
+lemma dist_pos_lt:
+  shows "x \<noteq> y ==> 0 < dist x y"
+by (simp add: zero_less_dist_iff)
+
+lemma dist_nz:
+  shows "x \<noteq> y \<longleftrightarrow> 0 < dist x y"
+by (simp add: zero_less_dist_iff)
+
+lemma dist_triangle_le:
+  shows "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e"
+by (rule order_trans [OF dist_triangle2])
+
+lemma dist_triangle_lt:
+  shows "dist x z + dist y z < e ==> dist x y < e"
+by (rule le_less_trans [OF dist_triangle2])
+
+lemma dist_triangle_half_l:
+  shows "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 \<Longrightarrow> dist x1 x2 < e"
+by (rule dist_triangle_lt [where z=y], simp)
+
+lemma dist_triangle_half_r:
+  shows "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 \<Longrightarrow> dist x1 x2 < e"
+by (rule dist_triangle_half_l, simp_all add: dist_commute)
+
 subclass topological_space
 proof
   have "\<exists>e::real. 0 < e"
@@ -554,6 +582,13 @@
     unfolding open_dist by fast
 qed
 
+lemma (in metric_space) open_ball: "open {y. dist x y < d}"
+proof (unfold open_dist, intro ballI)
+  fix y assume *: "y \<in> {y. dist x y < d}"
+  then show "\<exists>e>0. \<forall>z. dist z y < e \<longrightarrow> z \<in> {y. dist x y < d}"
+    by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt)
+qed
+
 end
 
 
@@ -1060,4 +1095,78 @@
 interpretation of_real: bounded_linear "\<lambda>r. of_real r"
 unfolding of_real_def by (rule scaleR.bounded_linear_left)
 
+subsection{* Hausdorff and other separation properties *}
+
+class t0_space = topological_space +
+  assumes t0_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> \<not> (x \<in> U \<longleftrightarrow> y \<in> U)"
+
+class t1_space = topological_space +
+  assumes t1_space: "x \<noteq> y \<Longrightarrow> \<exists>U. open U \<and> x \<in> U \<and> y \<notin> U"
+
+instance t1_space \<subseteq> t0_space
+proof qed (fast dest: t1_space)
+
+lemma separation_t1:
+  fixes x y :: "'a::t1_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> x \<in> U \<and> y \<notin> U)"
+  using t1_space[of x y] by blast
+
+lemma closed_singleton:
+  fixes a :: "'a::t1_space"
+  shows "closed {a}"
+proof -
+  let ?T = "\<Union>{S. open S \<and> a \<notin> S}"
+  have "open ?T" by (simp add: open_Union)
+  also have "?T = - {a}"
+    by (simp add: set_eq_iff separation_t1, auto)
+  finally show "closed {a}" unfolding closed_def .
+qed
+
+lemma closed_insert [simp]:
+  fixes a :: "'a::t1_space"
+  assumes "closed S" shows "closed (insert a S)"
+proof -
+  from closed_singleton assms
+  have "closed ({a} \<union> S)" by (rule closed_Un)
+  thus "closed (insert a S)" by simp
+qed
+
+lemma finite_imp_closed:
+  fixes S :: "'a::t1_space set"
+  shows "finite S \<Longrightarrow> closed S"
+by (induct set: finite, simp_all)
+
+text {* T2 spaces are also known as Hausdorff spaces. *}
+
+class t2_space = topological_space +
+  assumes hausdorff: "x \<noteq> y \<Longrightarrow> \<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+
+instance t2_space \<subseteq> t1_space
+proof qed (fast dest: hausdorff)
+
+instance metric_space \<subseteq> t2_space
+proof
+  fix x y :: "'a::metric_space"
+  assume xy: "x \<noteq> y"
+  let ?U = "{y'. dist x y' < dist x y / 2}"
+  let ?V = "{x'. dist y x' < dist x y / 2}"
+  have th0: "\<And>d x y z. (d x z :: real) \<le> d x y + d y z \<Longrightarrow> d y z = d z y
+               \<Longrightarrow> \<not>(d x y * 2 < d x z \<and> d z y * 2 < d x z)" by arith
+  have "open ?U \<and> open ?V \<and> x \<in> ?U \<and> y \<in> ?V \<and> ?U \<inter> ?V = {}"
+    using dist_pos_lt[OF xy] th0[of dist, OF dist_triangle dist_commute]
+    using open_ball[of _ "dist x y / 2"] by auto
+  then show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+    by blast
+qed
+
+lemma separation_t2:
+  fixes x y :: "'a::t2_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
+  using hausdorff[of x y] by blast
+
+lemma separation_t0:
+  fixes x y :: "'a::t0_space"
+  shows "x \<noteq> y \<longleftrightarrow> (\<exists>U. open U \<and> ~(x\<in>U \<longleftrightarrow> y\<in>U))"
+  using t0_space[of x y] by blast
+
 end
--- a/src/HOL/SEQ.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/SEQ.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -13,25 +13,7 @@
 imports Limits RComplete
 begin
 
-abbreviation
-  LIMSEQ :: "[nat \<Rightarrow> 'a::topological_space, 'a] \<Rightarrow> bool"
-    ("((_)/ ----> (_))" [60, 60] 60) where
-  "X ----> L \<equiv> (X ---> L) sequentially"
-
-definition
-  lim :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a" where
-    --{*Standard definition of limit using choice operator*}
-  "lim X = (THE L. X ----> L)"
-
-definition
-  convergent :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
-    --{*Standard definition of convergence*}
-  "convergent X = (\<exists>L. X ----> L)"
-
-definition
-  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
-    --{*Standard definition for bounded sequence*}
-  "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
+subsection {* Monotone sequences and subsequences *}
 
 definition
   monoseq :: "(nat \<Rightarrow> 'a::order) \<Rightarrow> bool" where
@@ -56,6 +38,171 @@
     --{*Definition of subsequence*}
   "subseq f \<longleftrightarrow> (\<forall>m. \<forall>n>m. f m < f n)"
 
+lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
+  unfolding mono_def incseq_def by auto
+
+lemma incseq_SucI:
+  "(\<And>n. X n \<le> X (Suc n)) \<Longrightarrow> incseq X"
+  using lift_Suc_mono_le[of X]
+  by (auto simp: incseq_def)
+
+lemma incseqD: "\<And>i j. incseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f i \<le> f j"
+  by (auto simp: incseq_def)
+
+lemma incseq_SucD: "incseq A \<Longrightarrow> A i \<le> A (Suc i)"
+  using incseqD[of A i "Suc i"] by auto
+
+lemma incseq_Suc_iff: "incseq f \<longleftrightarrow> (\<forall>n. f n \<le> f (Suc n))"
+  by (auto intro: incseq_SucI dest: incseq_SucD)
+
+lemma incseq_const[simp, intro]: "incseq (\<lambda>x. k)"
+  unfolding incseq_def by auto
+
+lemma decseq_SucI:
+  "(\<And>n. X (Suc n) \<le> X n) \<Longrightarrow> decseq X"
+  using order.lift_Suc_mono_le[OF dual_order, of X]
+  by (auto simp: decseq_def)
+
+lemma decseqD: "\<And>i j. decseq f \<Longrightarrow> i \<le> j \<Longrightarrow> f j \<le> f i"
+  by (auto simp: decseq_def)
+
+lemma decseq_SucD: "decseq A \<Longrightarrow> A (Suc i) \<le> A i"
+  using decseqD[of A i "Suc i"] by auto
+
+lemma decseq_Suc_iff: "decseq f \<longleftrightarrow> (\<forall>n. f (Suc n) \<le> f n)"
+  by (auto intro: decseq_SucI dest: decseq_SucD)
+
+lemma decseq_const[simp, intro]: "decseq (\<lambda>x. k)"
+  unfolding decseq_def by auto
+
+lemma monoseq_iff: "monoseq X \<longleftrightarrow> incseq X \<or> decseq X"
+  unfolding monoseq_def incseq_def decseq_def ..
+
+lemma monoseq_Suc:
+  "monoseq X \<longleftrightarrow> (\<forall>n. X n \<le> X (Suc n)) \<or> (\<forall>n. X (Suc n) \<le> X n)"
+  unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff ..
+
+lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
+by (simp add: monoseq_def)
+
+lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
+by (simp add: monoseq_Suc)
+
+lemma monoseq_minus:
+  fixes a :: "nat \<Rightarrow> 'a::ordered_ab_group_add"
+  assumes "monoseq a"
+  shows "monoseq (\<lambda> n. - a n)"
+proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
+  case True
+  hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto
+  thus ?thesis by (rule monoI2)
+next
+  case False
+  hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
+  thus ?thesis by (rule monoI1)
+qed
+
+text{*Subsequence (alternative definition, (e.g. Hoskins)*}
+
+lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
+apply (simp add: subseq_def)
+apply (auto dest!: less_imp_Suc_add)
+apply (induct_tac k)
+apply (auto intro: less_trans)
+done
+
+text{* for any sequence, there is a monotonic subsequence *}
+lemma seq_monosub:
+  fixes s :: "nat => 'a::linorder"
+  shows "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
+proof cases
+  let "?P p n" = "p > n \<and> (\<forall>m\<ge>p. s m \<le> s p)"
+  assume *: "\<forall>n. \<exists>p. ?P p n"
+  def f \<equiv> "nat_rec (SOME p. ?P p 0) (\<lambda>_ n. SOME p. ?P p n)"
+  have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp
+  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
+  have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto
+  have P_Suc: "\<And>i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto
+  then have "subseq f" unfolding subseq_Suc_iff by auto
+  moreover have "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc
+  proof (intro disjI2 allI)
+    fix n show "s (f (Suc n)) \<le> s (f n)"
+    proof (cases n)
+      case 0 with P_Suc[of 0] P_0 show ?thesis by auto
+    next
+      case (Suc m)
+      from P_Suc[of n] Suc have "f (Suc m) \<le> f (Suc (Suc m))" by simp
+      with P_Suc Suc show ?thesis by simp
+    qed
+  qed
+  ultimately show ?thesis by auto
+next
+  let "?P p m" = "m < p \<and> s m < s p"
+  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. s m \<le> s p))"
+  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. s p < s m" by (force simp: not_le le_less)
+  def f \<equiv> "nat_rec (SOME p. ?P p (Suc N)) (\<lambda>_ n. SOME p. ?P p n)"
+  have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp
+  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
+  have P_0: "?P (f 0) (Suc N)"
+    unfolding f_0 some_eq_ex[of "\<lambda>p. ?P p (Suc N)"] using N[of "Suc N"] by auto
+  { fix i have "N < f i \<Longrightarrow> ?P (f (Suc i)) (f i)"
+      unfolding f_Suc some_eq_ex[of "\<lambda>p. ?P p (f i)"] using N[of "f i"] . }
+  note P' = this
+  { fix i have "N < f i \<and> ?P (f (Suc i)) (f i)"
+      by (induct i) (insert P_0 P', auto) }
+  then have "subseq f" "monoseq (\<lambda>x. s (f x))"
+    unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le)
+  then show ?thesis by auto
+qed
+
+lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
+  have "n < f (Suc n)" by arith
+  thus ?case by arith
+qed
+
+lemma incseq_imp_monoseq:  "incseq X \<Longrightarrow> monoseq X"
+  by (simp add: incseq_def monoseq_def)
+
+lemma decseq_imp_monoseq:  "decseq X \<Longrightarrow> monoseq X"
+  by (simp add: decseq_def monoseq_def)
+
+lemma decseq_eq_incseq:
+  fixes X :: "nat \<Rightarrow> 'a::ordered_ab_group_add" shows "decseq X = incseq (\<lambda>n. - X n)" 
+  by (simp add: decseq_def incseq_def)
+
+subsection {* Defintions of limits *}
+
+abbreviation
+  LIMSEQ :: "[nat \<Rightarrow> 'a::topological_space, 'a] \<Rightarrow> bool"
+    ("((_)/ ----> (_))" [60, 60] 60) where
+  "X ----> L \<equiv> (X ---> L) sequentially"
+
+definition
+  lim :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a" where
+    --{*Standard definition of limit using choice operator*}
+  "lim X = (THE L. X ----> L)"
+
+definition
+  convergent :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
+    --{*Standard definition of convergence*}
+  "convergent X = (\<exists>L. X ----> L)"
+
+definition
+  Bseq :: "(nat => 'a::real_normed_vector) => bool" where
+    --{*Standard definition for bounded sequence*}
+  "Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
+
 definition
   Cauchy :: "(nat \<Rightarrow> 'a::metric_space) \<Rightarrow> bool" where
     --{*Standard definition of the Cauchy condition*}
@@ -502,53 +649,6 @@
     qed
 qed
 
-text{*Subsequence (alternative definition, (e.g. Hoskins)*}
-
-lemma subseq_Suc_iff: "subseq f = (\<forall>n. (f n) < (f (Suc n)))"
-apply (simp add: subseq_def)
-apply (auto dest!: less_imp_Suc_add)
-apply (induct_tac k)
-apply (auto intro: less_trans)
-done
-
-lemma monoseq_Suc:
-  "monoseq X \<longleftrightarrow> (\<forall>n. X n \<le> X (Suc n)) \<or> (\<forall>n. X (Suc n) \<le> X n)"
-apply (simp add: monoseq_def)
-apply (auto dest!: le_imp_less_or_eq)
-apply (auto intro!: lessI [THEN less_imp_le] dest!: less_imp_Suc_add)
-apply (induct_tac "ka")
-apply (auto intro: order_trans)
-apply (erule contrapos_np)
-apply (induct_tac "k")
-apply (auto intro: order_trans)
-done
-
-lemma monoI1: "\<forall>m. \<forall> n \<ge> m. X m \<le> X n ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma monoI2: "\<forall>m. \<forall> n \<ge> m. X n \<le> X m ==> monoseq X"
-by (simp add: monoseq_def)
-
-lemma mono_SucI1: "\<forall>n. X n \<le> X (Suc n) ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-lemma mono_SucI2: "\<forall>n. X (Suc n) \<le> X n ==> monoseq X"
-by (simp add: monoseq_Suc)
-
-lemma monoseq_minus:
-  fixes a :: "nat \<Rightarrow> 'a::ordered_ab_group_add"
-  assumes "monoseq a"
-  shows "monoseq (\<lambda> n. - a n)"
-proof (cases "\<forall> m. \<forall> n \<ge> m. a m \<le> a n")
-  case True
-  hence "\<forall> m. \<forall> n \<ge> m. - a n \<le> - a m" by auto
-  thus ?thesis by (rule monoI2)
-next
-  case False
-  hence "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" using `monoseq a`[unfolded monoseq_def] by auto
-  thus ?thesis by (rule monoI1)
-qed
-
 lemma monoseq_le:
   fixes a :: "nat \<Rightarrow> real"
   assumes "monoseq a" and "a ----> x"
@@ -602,60 +702,6 @@
   qed auto
 qed
 
-text{* for any sequence, there is a monotonic subsequence *}
-lemma seq_monosub:
-  fixes s :: "nat => 'a::linorder"
-  shows "\<exists>f. subseq f \<and> monoseq (\<lambda> n. (s (f n)))"
-proof cases
-  let "?P p n" = "p > n \<and> (\<forall>m\<ge>p. s m \<le> s p)"
-  assume *: "\<forall>n. \<exists>p. ?P p n"
-  def f \<equiv> "nat_rec (SOME p. ?P p 0) (\<lambda>_ n. SOME p. ?P p n)"
-  have f_0: "f 0 = (SOME p. ?P p 0)" unfolding f_def by simp
-  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
-  have P_0: "?P (f 0) 0" unfolding f_0 using *[rule_format] by (rule someI2_ex) auto
-  have P_Suc: "\<And>i. ?P (f (Suc i)) (f i)" unfolding f_Suc using *[rule_format] by (rule someI2_ex) auto
-  then have "subseq f" unfolding subseq_Suc_iff by auto
-  moreover have "monoseq (\<lambda>n. s (f n))" unfolding monoseq_Suc
-  proof (intro disjI2 allI)
-    fix n show "s (f (Suc n)) \<le> s (f n)"
-    proof (cases n)
-      case 0 with P_Suc[of 0] P_0 show ?thesis by auto
-    next
-      case (Suc m)
-      from P_Suc[of n] Suc have "f (Suc m) \<le> f (Suc (Suc m))" by simp
-      with P_Suc Suc show ?thesis by simp
-    qed
-  qed
-  ultimately show ?thesis by auto
-next
-  let "?P p m" = "m < p \<and> s m < s p"
-  assume "\<not> (\<forall>n. \<exists>p>n. (\<forall>m\<ge>p. s m \<le> s p))"
-  then obtain N where N: "\<And>p. p > N \<Longrightarrow> \<exists>m>p. s p < s m" by (force simp: not_le le_less)
-  def f \<equiv> "nat_rec (SOME p. ?P p (Suc N)) (\<lambda>_ n. SOME p. ?P p n)"
-  have f_0: "f 0 = (SOME p. ?P p (Suc N))" unfolding f_def by simp
-  have f_Suc: "\<And>i. f (Suc i) = (SOME p. ?P p (f i))" unfolding f_def nat_rec_Suc ..
-  have P_0: "?P (f 0) (Suc N)"
-    unfolding f_0 some_eq_ex[of "\<lambda>p. ?P p (Suc N)"] using N[of "Suc N"] by auto
-  { fix i have "N < f i \<Longrightarrow> ?P (f (Suc i)) (f i)"
-      unfolding f_Suc some_eq_ex[of "\<lambda>p. ?P p (f i)"] using N[of "f i"] . }
-  note P' = this
-  { fix i have "N < f i \<and> ?P (f (Suc i)) (f i)"
-      by (induct i) (insert P_0 P', auto) }
-  then have "subseq f" "monoseq (\<lambda>x. s (f x))"
-    unfolding subseq_Suc_iff monoseq_Suc by (auto simp: not_le intro: less_imp_le)
-  then show ?thesis by auto
-qed
-
-lemma seq_suble: assumes sf: "subseq f" shows "n \<le> f n"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  from sf[unfolded subseq_Suc_iff, rule_format, of n] Suc.hyps
-  have "n < f (Suc n)" by arith
-  thus ?case by arith
-qed
-
 lemma LIMSEQ_subseq_LIMSEQ:
   "\<lbrakk> X ----> L; subseq f \<rbrakk> \<Longrightarrow> (X o f) ----> L"
 apply (rule topological_tendstoI)
@@ -810,9 +856,6 @@
 
 subsubsection{*Increasing and Decreasing Series*}
 
-lemma incseq_imp_monoseq:  "incseq X \<Longrightarrow> monoseq X"
-  by (simp add: incseq_def monoseq_def)
-
 lemma incseq_le:
   fixes X :: "nat \<Rightarrow> real"
   assumes inc: "incseq X" and lim: "X ----> L" shows "X n \<le> L"
@@ -834,32 +877,6 @@
     by (blast intro: eq_refl X)
 qed
 
-lemma incseq_SucI:
-  assumes "\<And>n. X n \<le> X (Suc n)"
-  shows "incseq X" unfolding incseq_def
-proof safe
-  fix m n :: nat
-  { fix d m :: nat
-    have "X m \<le> X (m + d)"
-    proof (induct d)
-      case (Suc d)
-      also have "X (m + d) \<le> X (m + Suc d)"
-        using assms by simp
-      finally show ?case .
-    qed simp }
-  note this[of m "n - m"]
-  moreover assume "m \<le> n"
-  ultimately show "X m \<le> X n" by simp
-qed
-
-lemma decseq_imp_monoseq:  "decseq X \<Longrightarrow> monoseq X"
-  by (simp add: decseq_def monoseq_def)
-
-lemma decseq_eq_incseq:
-  fixes X :: "nat \<Rightarrow> 'a::ordered_ab_group_add" shows "decseq X = incseq (\<lambda>n. - X n)" 
-  by (simp add: decseq_def incseq_def)
-
-
 lemma decseq_le:
   fixes X :: "nat \<Rightarrow> real" assumes dec: "decseq X" and lim: "X ----> L" shows "L \<le> X n"
 proof -
--- a/src/HOL/Series.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Series.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -5,7 +5,7 @@
 Converted to Isar and polished by lcp
 Converted to setsum and polished yet more by TNN
 Additional contributions by Jeremy Avigad
-*) 
+*)
 
 header{*Finite Summation and Infinite Series*}
 
@@ -14,16 +14,16 @@
 begin
 
 definition
-   sums  :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
+   sums  :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a \<Rightarrow> bool"
      (infixr "sums" 80) where
    "f sums s = (%n. setsum f {0..<n}) ----> s"
 
 definition
-   summable :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool" where
+   summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> bool" where
    "summable f = (\<exists>s. f sums s)"
 
 definition
-   suminf   :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a" where
+   suminf   :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a" where
    "suminf f = (THE s. f sums s)"
 
 syntax
@@ -81,62 +81,65 @@
   "\<forall>n f. setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)::real) + setsum f {0..<k}"
 by (clarify, rule sumr_offset3)
 
-(*
-lemma sumr_from_1_from_0: "0 < n ==>
-      (\<Sum>n=Suc 0 ..< Suc n. if even(n) then 0 else
-             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n =
-      (\<Sum>n=0..<Suc n. if even(n) then 0 else
-             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n"
-by (rule_tac n1 = 1 in sumr_split_add [THEN subst], auto)
-*)
-
 subsection{* Infinite Sums, by the Properties of Limits*}
 
 (*----------------------
-   suminf is the sum   
+   suminf is the sum
  ---------------------*)
 lemma sums_summable: "f sums l ==> summable f"
-by (simp add: sums_def summable_def, blast)
+  by (simp add: sums_def summable_def, blast)
 
-lemma summable_sums: "summable f ==> f sums (suminf f)"
-apply (simp add: summable_def suminf_def sums_def)
-apply (fast intro: theI LIMSEQ_unique)
-done
+lemma summable_sums:
+  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}" assumes "summable f" shows "f sums (suminf f)"
+proof -
+  from assms guess s unfolding summable_def sums_def_raw .. note s = this
+  then show ?thesis unfolding sums_def_raw suminf_def
+    by (rule theI, auto intro!: tendsto_unique[OF trivial_limit_sequentially])
+qed
 
-lemma summable_sumr_LIMSEQ_suminf: 
-     "summable f ==> (%n. setsum f {0..<n}) ----> (suminf f)"
+lemma summable_sumr_LIMSEQ_suminf:
+  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
+  shows "summable f \<Longrightarrow> (\<lambda>n. setsum f {0..<n}) ----> suminf f"
 by (rule summable_sums [unfolded sums_def])
 
 lemma suminf_eq_lim: "suminf f = lim (%n. setsum f {0..<n})"
-  by (simp add: suminf_def sums_def lim_def) 
+  by (simp add: suminf_def sums_def lim_def)
 
 (*-------------------
-    sum is unique                    
+    sum is unique
  ------------------*)
-lemma sums_unique: "f sums s ==> (s = suminf f)"
-apply (frule sums_summable [THEN summable_sums])
-apply (auto intro!: LIMSEQ_unique simp add: sums_def)
+lemma sums_unique:
+  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
+  shows "f sums s \<Longrightarrow> (s = suminf f)"
+apply (frule sums_summable[THEN summable_sums])
+apply (auto intro!: tendsto_unique[OF trivial_limit_sequentially] simp add: sums_def)
 done
 
-lemma sums_iff: "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"
+lemma sums_iff:
+  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
+  shows "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"
   by (metis summable_sums sums_summable sums_unique)
 
-lemma sums_split_initial_segment: "f sums s ==> 
-  (%n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"
-  apply (unfold sums_def);
-  apply (simp add: sumr_offset); 
+lemma sums_split_initial_segment:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  shows "f sums s ==> (\<lambda>n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"
+  apply (unfold sums_def)
+  apply (simp add: sumr_offset)
   apply (rule LIMSEQ_diff_const)
   apply (rule LIMSEQ_ignore_initial_segment)
   apply assumption
 done
 
-lemma summable_ignore_initial_segment: "summable f ==> 
-    summable (%n. f(n + k))"
+lemma summable_ignore_initial_segment:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  shows "summable f ==> summable (%n. f(n + k))"
   apply (unfold summable_def)
   apply (auto intro: sums_split_initial_segment)
 done
 
-lemma suminf_minus_initial_segment: "summable f ==>
+lemma suminf_minus_initial_segment:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  shows "summable f ==>
     suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)"
   apply (frule summable_ignore_initial_segment)
   apply (rule sums_unique [THEN sym])
@@ -145,8 +148,10 @@
   apply auto
 done
 
-lemma suminf_split_initial_segment: "summable f ==> 
-    suminf f = (SUM i = 0..< k. f i) + suminf (%n. f(n + k))"
+lemma suminf_split_initial_segment:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  shows "summable f ==>
+    suminf f = (SUM i = 0..< k. f i) + (\<Sum>n. f(n + k))"
 by (auto simp add: suminf_minus_initial_segment)
 
 lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable a"
@@ -158,31 +163,42 @@
     by auto
 qed
 
-lemma sums_Suc: assumes sumSuc: "(\<lambda> n. f (Suc n)) sums l" shows "f sums (l + f 0)"
+lemma sums_Suc:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+  assumes sumSuc: "(\<lambda> n. f (Suc n)) sums l" shows "f sums (l + f 0)"
 proof -
   from sumSuc[unfolded sums_def]
   have "(\<lambda>i. \<Sum>n = Suc 0..<Suc i. f n) ----> l" unfolding setsum_reindex[OF inj_Suc] image_Suc_atLeastLessThan[symmetric] comp_def .
-  from LIMSEQ_add_const[OF this, where b="f 0"] 
+  from LIMSEQ_add_const[OF this, where b="f 0"]
   have "(\<lambda>i. \<Sum>n = 0..<Suc i. f n) ----> l + f 0" unfolding add_commute setsum_head_upt_Suc[OF zero_less_Suc] .
   thus ?thesis unfolding sums_def by (rule LIMSEQ_imp_Suc)
 qed
 
-lemma series_zero: 
-     "(\<forall>m. n \<le> m --> f(m) = 0) ==> f sums (setsum f {0..<n})"
-apply (simp add: sums_def LIMSEQ_iff diff_minus[symmetric], safe)
-apply (rule_tac x = n in exI)
-apply (clarsimp simp add:setsum_diff[symmetric] cong:setsum_ivl_cong)
-done
+lemma series_zero:
+  fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
+  assumes "\<forall>m. n \<le> m \<longrightarrow> f m = 0"
+  shows "f sums (setsum f {0..<n})"
+proof -
+  { fix k :: nat have "setsum f {0..<k + n} = setsum f {0..<n}"
+      using assms by (induct k) auto }
+  note setsum_const = this
+  show ?thesis
+    unfolding sums_def
+    apply (rule LIMSEQ_offset[of _ n])
+    unfolding setsum_const
+    apply (rule LIMSEQ_const)
+    done
+qed
 
-lemma sums_zero: "(\<lambda>n. 0) sums 0"
+lemma sums_zero[simp, intro]: "(\<lambda>n. 0) sums 0"
 unfolding sums_def by (simp add: LIMSEQ_const)
 
-lemma summable_zero: "summable (\<lambda>n. 0)"
+lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"
 by (rule sums_zero [THEN sums_summable])
 
-lemma suminf_zero: "suminf (\<lambda>n. 0) = 0"
+lemma suminf_zero[simp]: "suminf (\<lambda>n. 0::'a::{t2_space, comm_monoid_add}) = 0"
 by (rule sums_zero [THEN sums_unique, symmetric])
-  
+
 lemma (in bounded_linear) sums:
   "(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"
 unfolding sums_def by (drule LIMSEQ, simp only: setsum)
@@ -207,7 +223,7 @@
 
 lemma suminf_mult:
   fixes c :: "'a::real_normed_algebra"
-  shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f";
+  shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f"
 by (rule mult_right.suminf [symmetric])
 
 lemma sums_mult2:
@@ -240,37 +256,54 @@
   shows "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"
 by (rule divide.suminf [symmetric])
 
-lemma sums_add: "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) sums (a + b)"
+lemma sums_add:
+  fixes a b :: "'a::real_normed_field"
+  shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) sums (a + b)"
 unfolding sums_def by (simp add: setsum_addf LIMSEQ_add)
 
-lemma summable_add: "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"
+lemma summable_add:
+  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"
 unfolding summable_def by (auto intro: sums_add)
 
 lemma suminf_add:
-  "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"
+  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"
 by (intro sums_unique sums_add summable_sums)
 
-lemma sums_diff: "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) sums (a - b)"
+lemma sums_diff:
+  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) sums (a - b)"
 unfolding sums_def by (simp add: setsum_subtractf LIMSEQ_diff)
 
-lemma summable_diff: "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"
+lemma summable_diff:
+  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"
 unfolding summable_def by (auto intro: sums_diff)
 
 lemma suminf_diff:
-  "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"
+  fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"
 by (intro sums_unique sums_diff summable_sums)
 
-lemma sums_minus: "X sums a ==> (\<lambda>n. - X n) sums (- a)"
+lemma sums_minus:
+  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "X sums a ==> (\<lambda>n. - X n) sums (- a)"
 unfolding sums_def by (simp add: setsum_negf LIMSEQ_minus)
 
-lemma summable_minus: "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"
+lemma summable_minus:
+  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"
 unfolding summable_def by (auto intro: sums_minus)
 
-lemma suminf_minus: "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"
+lemma suminf_minus:
+  fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"
 by (intro sums_unique [symmetric] sums_minus summable_sums)
 
 lemma sums_group:
-     "[|summable f; 0 < k |] ==> (%n. setsum f {n*k..<n*k+k}) sums (suminf f)"
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "[|summable f; 0 < k |] ==> (%n. setsum f {n*k..<n*k+k}) sums (suminf f)"
 apply (drule summable_sums)
 apply (simp only: sums_def sumr_group)
 apply (unfold LIMSEQ_iff, safe)
@@ -290,19 +323,19 @@
   assumes pos: "!!n. 0 \<le> f n" and le: "!!n. setsum f {0..<n} \<le> x"
   shows "summable f"
 proof -
-  have "convergent (\<lambda>n. setsum f {0..<n})" 
+  have "convergent (\<lambda>n. setsum f {0..<n})"
     proof (rule Bseq_mono_convergent)
       show "Bseq (\<lambda>n. setsum f {0..<n})"
         by (rule f_inc_g_dec_Beq_f [of "(\<lambda>n. setsum f {0..<n})" "\<lambda>n. x"])
-           (auto simp add: le pos)  
-    next 
+           (auto simp add: le pos)
+    next
       show "\<forall>m n. m \<le> n \<longrightarrow> setsum f {0..<m} \<le> setsum f {0..<n}"
-        by (auto intro: setsum_mono2 pos) 
+        by (auto intro: setsum_mono2 pos)
     qed
   then obtain L where "(%n. setsum f {0..<n}) ----> L"
     by (blast dest: convergentD)
   thus ?thesis
-    by (force simp add: summable_def sums_def) 
+    by (force simp add: summable_def sums_def)
 qed
 
 lemma series_pos_le:
@@ -382,7 +415,7 @@
 by (rule geometric_sums [THEN sums_summable])
 
 lemma half: "0 < 1 / (2::'a::{number_ring,linordered_field_inverse_zero})"
-  by arith 
+  by arith
 
 lemma power_half_series: "(\<lambda>n. (1/2::real)^Suc n) sums 1"
 proof -
@@ -400,7 +433,9 @@
  "summable f = convergent (%n. setsum f {0..<n})"
 by (simp add: summable_def sums_def convergent_def)
 
-lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"
+lemma summable_LIMSEQ_zero:
+  fixes f :: "nat \<Rightarrow> 'a::real_normed_field"
+  shows "summable f \<Longrightarrow> f ----> 0"
 apply (drule summable_convergent_sumr_iff [THEN iffD1])
 apply (drule convergent_Cauchy)
 apply (simp only: Cauchy_iff LIMSEQ_iff, safe)
@@ -413,10 +448,10 @@
 lemma suminf_le:
   fixes x :: real
   shows "summable f \<Longrightarrow> (!!n. setsum f {0..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
-  by (simp add: summable_convergent_sumr_iff suminf_eq_lim lim_le) 
+  by (simp add: summable_convergent_sumr_iff suminf_eq_lim lim_le)
 
 lemma summable_Cauchy:
-     "summable (f::nat \<Rightarrow> 'a::banach) =  
+     "summable (f::nat \<Rightarrow> 'a::banach) =
       (\<forall>e > 0. \<exists>N. \<forall>m \<ge> N. \<forall>n. norm (setsum f {m..<n}) < e)"
 apply (simp only: summable_convergent_sumr_iff Cauchy_convergent_iff [symmetric] Cauchy_iff, safe)
 apply (drule spec, drule (1) mp)
@@ -522,7 +557,7 @@
   moreover from sm have "summable f" .
   ultimately have "suminf ?g \<le> suminf f" by (rule summable_le)
   then show "0 \<le> suminf f" by (simp add: suminf_zero)
-qed 
+qed
 
 
 text{*Absolute convergence imples normal convergence*}
@@ -596,7 +631,7 @@
   fixes f :: "nat \<Rightarrow> 'a::banach"
   shows "\<lbrakk>c < 1; \<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> summable f"
 apply (frule ratio_test_lemma2, auto)
-apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n" 
+apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n"
        in summable_comparison_test)
 apply (rule_tac x = N in exI, safe)
 apply (drule le_Suc_ex_iff [THEN iffD1])
@@ -605,7 +640,7 @@
 apply (rule_tac y = "c * norm (f (N + n))" in order_trans)
 apply (auto intro: mult_right_mono simp add: summable_def)
 apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)
-apply (rule sums_divide) 
+apply (rule sums_divide)
 apply (rule sums_mult)
 apply (auto intro!: geometric_sums)
 done
--- a/src/HOL/Transcendental.thy	Mon Mar 14 15:17:10 2011 +0100
+++ b/src/HOL/Transcendental.thy	Mon Mar 14 15:29:10 2011 +0100
@@ -22,14 +22,14 @@
 
 lemma lemma_realpow_diff_sumr:
   fixes y :: "'a::{comm_semiring_0,monoid_mult}" shows
-     "(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) =  
+     "(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) =
       y * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
 by (simp add: setsum_right_distrib lemma_realpow_diff mult_ac
          del: setsum_op_ivl_Suc)
 
 lemma lemma_realpow_diff_sumr2:
   fixes y :: "'a::{comm_ring,monoid_mult}" shows
-     "x ^ (Suc n) - y ^ (Suc n) =  
+     "x ^ (Suc n) - y ^ (Suc n) =
       (x - y) * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
 apply (induct n, simp)
 apply (simp del: setsum_op_ivl_Suc)
@@ -42,7 +42,7 @@
 done
 
 lemma lemma_realpow_rev_sumr:
-     "(\<Sum>p=0..<Suc n. (x ^ p) * (y ^ (n - p))) =  
+     "(\<Sum>p=0..<Suc n. (x ^ p) * (y ^ (n - p))) =
       (\<Sum>p=0..<Suc n. (x ^ (n - p)) * (y ^ p))"
 apply (rule setsum_reindex_cong [where f="\<lambda>i. n - i"])
 apply (rule inj_onI, simp)
@@ -107,16 +107,16 @@
 
 lemma powser_inside:
   fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach}" shows
-     "[| summable (%n. f(n) * (x ^ n)); norm z < norm x |]  
+     "[| summable (%n. f(n) * (x ^ n)); norm z < norm x |]
       ==> summable (%n. f(n) * (z ^ n))"
 by (rule powser_insidea [THEN summable_norm_cancel])
 
 lemma sum_split_even_odd: fixes f :: "nat \<Rightarrow> real" shows
-  "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) = 
+  "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) =
    (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1))"
 proof (induct n)
   case (Suc n)
-  have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) = 
+  have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) =
         (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1)) + (f (2 * n) + g (2 * n + 1))"
     using Suc.hyps unfolding One_nat_def by auto
   also have "\<dots> = (\<Sum> i = 0 ..< Suc n. f (2 * i)) + (\<Sum> i = 0 ..< Suc n. g (2 * i + 1))" by auto
@@ -133,7 +133,7 @@
 
   let ?SUM = "\<lambda> m. \<Sum> i = 0 ..< m. if even i then 0 else g ((i - 1) div 2)"
   { fix m assume "m \<ge> 2 * no" hence "m div 2 \<ge> no" by auto
-    have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }" 
+    have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }"
       using sum_split_even_odd by auto
     hence "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using `m div 2 \<ge> no` by auto
     moreover
@@ -161,13 +161,13 @@
   { fix B T E have "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)"
       by (cases B) auto } note if_sum = this
   have g_sums: "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF `g sums x`] .
-  { 
+  {
     have "?s 0 = 0" by auto
     have Suc_m1: "\<And> n. Suc n - 1 = n" by auto
     have if_eq: "\<And>B T E. (if \<not> B then T else E) = (if B then E else T)" by auto
 
     have "?s sums y" using sums_if'[OF `f sums y`] .
-    from this[unfolded sums_def, THEN LIMSEQ_Suc] 
+    from this[unfolded sums_def, THEN LIMSEQ_Suc]
     have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
       unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
                 image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
@@ -181,7 +181,7 @@
 lemma sums_alternating_upper_lower:
   fixes a :: "nat \<Rightarrow> real"
   assumes mono: "\<And>n. a (Suc n) \<le> a n" and a_pos: "\<And>n. 0 \<le> a n" and "a ----> 0"
-  shows "\<exists>l. ((\<forall>n. (\<Sum>i=0..<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i=0..<2*n. -1^i*a i) ----> l) \<and> 
+  shows "\<exists>l. ((\<forall>n. (\<Sum>i=0..<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i=0..<2*n. -1^i*a i) ----> l) \<and>
              ((\<forall>n. l \<le> (\<Sum>i=0..<2*n + 1. -1^i*a i)) \<and> (\<lambda> n. \<Sum>i=0..<2*n + 1. -1^i*a i) ----> l)"
   (is "\<exists>l. ((\<forall>n. ?f n \<le> l) \<and> _) \<and> ((\<forall>n. l \<le> ?g n) \<and> _)")
 proof -
@@ -194,21 +194,21 @@
   proof fix n show "?g (Suc n) \<le> ?g n" using mono[of "Suc (2*n)"]
     unfolding One_nat_def by auto qed
   moreover
-  have "\<forall> n. ?f n \<le> ?g n" 
+  have "\<forall> n. ?f n \<le> ?g n"
   proof fix n show "?f n \<le> ?g n" using fg_diff a_pos
     unfolding One_nat_def by auto qed
   moreover
   have "(\<lambda> n. ?f n - ?g n) ----> 0" unfolding fg_diff
   proof (rule LIMSEQ_I)
     fix r :: real assume "0 < r"
-    with `a ----> 0`[THEN LIMSEQ_D] 
+    with `a ----> 0`[THEN LIMSEQ_D]
     obtain N where "\<And> n. n \<ge> N \<Longrightarrow> norm (a n - 0) < r" by auto
     hence "\<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
     thus "\<exists> N. \<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
   qed
   ultimately
   show ?thesis by (rule lemma_nest_unique)
-qed 
+qed
 
 lemma summable_Leibniz': fixes a :: "nat \<Rightarrow> real"
   assumes a_zero: "a ----> 0" and a_pos: "\<And> n. 0 \<le> a n"
@@ -225,16 +225,16 @@
   let "?g n" = "?P (2 * n + 1)"
   obtain l :: real where below_l: "\<forall> n. ?f n \<le> l" and "?f ----> l" and above_l: "\<forall> n. l \<le> ?g n" and "?g ----> l"
     using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast
-  
+
   let ?Sa = "\<lambda> m. \<Sum> n = 0..<m. ?S n"
   have "?Sa ----> l"
   proof (rule LIMSEQ_I)
     fix r :: real assume "0 < r"
 
-    with `?f ----> l`[THEN LIMSEQ_D] 
+    with `?f ----> l`[THEN LIMSEQ_D]
     obtain f_no where f: "\<And> n. n \<ge> f_no \<Longrightarrow> norm (?f n - l) < r" by auto
 
-    from `0 < r` `?g ----> l`[THEN LIMSEQ_D] 
+    from `0 < r` `?g ----> l`[THEN LIMSEQ_D]
     obtain g_no where g: "\<And> n. n \<ge> g_no \<Longrightarrow> norm (?g n - l) < r" by auto
 
     { fix n :: nat
@@ -302,7 +302,7 @@
     hence ?summable unfolding summable_def by auto
     moreover
     have "\<And> a b :: real. \<bar> - a - - b \<bar> = \<bar>a - b\<bar>" unfolding minus_diff_minus by auto
-    
+
     from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus]
     have move_minus: "(\<Sum>n. - (-1 ^ n * a n)) = - (\<Sum>n. -1 ^ n * a n)" by auto
 
@@ -336,8 +336,9 @@
 done
 
 lemma diffs_equiv:
-     "summable (%n. (diffs c)(n) * (x ^ n)) ==>  
-      (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums  
+  fixes x :: "'a::{real_normed_vector, ring_1}"
+  shows "summable (%n. (diffs c)(n) * (x ^ n)) ==>
+      (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums
          (\<Sum>n. (diffs c)(n) * (x ^ n))"
 unfolding diffs_def
 apply (drule summable_sums)
@@ -346,7 +347,7 @@
 
 lemma lemma_termdiff1:
   fixes z :: "'a :: {monoid_mult,comm_ring}" shows
-  "(\<Sum>p=0..<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =  
+  "(\<Sum>p=0..<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =
    (\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
 by(auto simp add: algebra_simps power_add [symmetric])
 
@@ -535,7 +536,7 @@
       apply (simp add: diffs_def)
       apply (case_tac n, simp_all add: r_neq_0)
       done
-    finally have "summable 
+    finally have "summable
       (\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))"
       by (rule diffs_equiv [THEN sums_summable])
     also have
@@ -596,7 +597,7 @@
     have C: "summable (\<lambda>n. diffs c n * x ^ n)"
       by (rule powser_inside [OF 2 4])
     show "((\<Sum>n. c n * (x + h) ^ n) - (\<Sum>n. c n * x ^ n)) / h
-             - (\<Sum>n. diffs c n * x ^ n) = 
+             - (\<Sum>n. diffs c n * x ^ n) =
           (\<Sum>n. c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))"
       apply (subst sums_unique [OF diffs_equiv [OF C]])
       apply (subst suminf_diff [OF B A])
@@ -646,10 +647,10 @@
 proof (rule LIM_I)
   fix r :: real assume "0 < r" hence "0 < r/3" by auto
 
-  obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3" 
+  obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3"
     using suminf_exist_split[OF `0 < r/3` `summable L`] by auto
 
-  obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3" 
+  obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3"
     using suminf_exist_split[OF `0 < r/3` `summable (f' x0)`] by auto
 
   let ?N = "Suc (max N_L N_f')"
@@ -672,7 +673,7 @@
     proof (rule ballI)
       fix x assume "x \<in> ?s ` {0..<?N}"
       then obtain n where "x = ?s n" and "n \<in> {0..<?N}" using image_iff[THEN iffD1] by blast
-      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def] 
+      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def]
       obtain s where s_bound: "0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < s \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r)" by auto
       have "0 < ?s n" by (rule someI2[where a=s], auto simp add: s_bound)
       thus "0 < x" unfolding `x = ?s n` .
@@ -685,7 +686,7 @@
 
   { fix x assume "x \<noteq> 0" and "\<bar> x \<bar> < S"
     hence x_in_I: "x0 + x \<in> { a <..< b }" using S_a S_b by auto
-    
+
     note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
     note div_smbl = summable_divide[OF diff_smbl]
     note all_smbl = summable_diff[OF div_smbl `summable (f' x0)`]
@@ -695,7 +696,7 @@
     note all_shft_smbl = summable_diff[OF div_smbl ign[OF `summable (f' x0)`]]
 
     { fix n
-      have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>" 
+      have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>"
         using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
       hence "\<bar> ( \<bar> ?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)" using `x \<noteq> 0` by auto
     } note L_ge = summable_le2[OF allI[OF this] ign[OF `summable L`]]
@@ -709,7 +710,7 @@
       fix n assume "n \<in> { 0 ..< ?N}"
       have "\<bar> x \<bar> < S" using `\<bar> x \<bar> < S` .
       also have "S \<le> S'" using `S \<le> S'` .
-      also have "S' \<le> ?s n" unfolding S'_def 
+      also have "S' \<le> ?s n" unfolding S'_def
       proof (rule Min_le_iff[THEN iffD2])
         have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
         thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
@@ -727,16 +728,16 @@
     finally have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> < r / 3" (is "?diff_part < r / 3") .
 
     from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
-    have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> = 
+    have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> =
                     \<bar> \<Sum>n. ?diff n x - f' x0 n \<bar>" unfolding suminf_diff[OF div_smbl `summable (f' x0)`, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto
     also have "\<dots> \<le> ?diff_part + \<bar> (\<Sum>n. ?diff (n + ?N) x) - (\<Sum> n. f' x0 (n + ?N)) \<bar>" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]] by (rule abs_triangle_ineq)
     also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto
-    also have "\<dots> < r /3 + r/3 + r/3" 
+    also have "\<dots> < r /3 + r/3 + r/3"
       using `?diff_part < r/3` `?L_part \<le> r/3` and `?f'_part < r/3`
       by (rule add_strict_mono [OF add_less_le_mono])
     finally have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> < r"
       by auto
-  } thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow> 
+  } thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow>
       norm (((\<Sum>n. f (x0 + x) n) - (\<Sum>n. f x0 n)) / x - (\<Sum>n. f' x0 n)) < r" using `0 < S`
     unfolding real_norm_def diff_0_right by blast
 qed
@@ -761,9 +762,9 @@
       { fix n x y assume "x \<in> {-R' <..< R'}" and "y \<in> {-R' <..< R'}"
         show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
         proof -
-          have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>" 
+          have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>"
             unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
-          also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)" 
+          also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)"
           proof (rule mult_left_mono)
             have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)" by (rule setsum_abs)
             also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
@@ -809,7 +810,7 @@
   next
     case False
     have "- ?R < 0" using assms by auto
-    also have "\<dots> \<le> x0" using False by auto 
+    also have "\<dots> \<le> x0" using False by auto
     finally show ?thesis .
   qed
   hence "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by auto
@@ -871,7 +872,7 @@
 unfolding exp_def by (rule summable_exp_generic [THEN summable_sums])
 
 
-lemma exp_fdiffs: 
+lemma exp_fdiffs:
       "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))"
 by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def of_nat_mult
          del: mult_Suc of_nat_Suc)
@@ -1081,7 +1082,7 @@
 lemma lemma_exp_total: "1 \<le> y ==> \<exists>x. 0 \<le> x & x \<le> y - 1 & exp(x::real) = y"
 apply (rule IVT)
 apply (auto intro: isCont_exp simp add: le_diff_eq)
-apply (subgoal_tac "1 + (y - 1) \<le> exp (y - 1)") 
+apply (subgoal_tac "1 + (y - 1) \<le> exp (y - 1)")
 apply simp
 apply (rule exp_ge_add_one_self_aux, simp)
 done
@@ -1160,12 +1161,12 @@
 qed
 
 lemma ln_ge_zero_imp_ge_one:
-  assumes ln: "0 \<le> ln x" 
+  assumes ln: "0 \<le> ln x"
       and x:  "0 < x"
   shows "1 \<le> x"
 proof -
   from ln have "ln 1 \<le> ln x" by simp
-  thus ?thesis by (simp add: x del: ln_one) 
+  thus ?thesis by (simp add: x del: ln_one)
 qed
 
 lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \<le> ln x) = (1 \<le> x)"
@@ -1183,12 +1184,12 @@
 qed
 
 lemma ln_gt_zero_imp_gt_one:
-  assumes ln: "0 < ln x" 
+  assumes ln: "0 < ln x"
       and x:  "0 < x"
   shows "1 < x"
 proof -
   from ln have "ln 1 < ln x" by simp
-  thus ?thesis by (simp add: x del: ln_one) 
+  thus ?thesis by (simp add: x del: ln_one)
 qed
 
 lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)"
@@ -1285,22 +1286,22 @@
 done
 
 lemma lemma_STAR_sin:
-     "(if even n then 0  
+     "(if even n then 0
        else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0"
 by (induct "n", auto)
 
 lemma lemma_STAR_cos:
-     "0 < n -->  
+     "0 < n -->
       -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"
 by (induct "n", auto)
 
 lemma lemma_STAR_cos1:
-     "0 < n -->  
+     "0 < n -->
       (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"
 by (induct "n", auto)
 
 lemma lemma_STAR_cos2:
-  "(\<Sum>n=1..<n. if even n then -1 ^ (n div 2)/(real (fact n)) *  0 ^ n 
+  "(\<Sum>n=1..<n. if even n then -1 ^ (n div 2)/(real (fact n)) *  0 ^ n
                          else 0) = 0"
 apply (induct "n")
 apply (case_tac [2] "n", auto)
@@ -1314,7 +1315,7 @@
 
 lemma sin_fdiffs: "diffs sin_coeff = cos_coeff"
 unfolding sin_coeff_def cos_coeff_def
-by (auto intro!: ext 
+by (auto intro!: ext
          simp add: diffs_def divide_inverse real_of_nat_def of_nat_mult
          simp del: mult_Suc of_nat_Suc)
 
@@ -1323,7 +1324,7 @@
 
 lemma cos_fdiffs: "diffs cos_coeff = (\<lambda>n. - sin_coeff n)"
 unfolding sin_coeff_def cos_coeff_def
-by (auto intro!: ext 
+by (auto intro!: ext
          simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def of_nat_mult
          simp del: mult_Suc of_nat_Suc)
 
@@ -1424,8 +1425,8 @@
      "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))"
   by (auto intro!: DERIV_intros)
 
-lemma DERIV_sin_circle_all: 
-     "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :>  
+lemma DERIV_sin_circle_all:
+     "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :>
              (2*cos(x)*sin(x) - 2*cos(x)*sin(x))"
   by (auto intro!: DERIV_intros)
 
@@ -1462,29 +1463,29 @@
 by (rule power2_le_imp_le, simp_all add: sin_squared_eq)
 
 lemma sin_ge_minus_one [simp]: "-1 \<le> sin x"
-apply (insert abs_sin_le_one [of x]) 
-apply (simp add: abs_le_iff del: abs_sin_le_one) 
+apply (insert abs_sin_le_one [of x])
+apply (simp add: abs_le_iff del: abs_sin_le_one)
 done
 
 lemma sin_le_one [simp]: "sin x \<le> 1"
-apply (insert abs_sin_le_one [of x]) 
-apply (simp add: abs_le_iff del: abs_sin_le_one) 
+apply (insert abs_sin_le_one [of x])
+apply (simp add: abs_le_iff del: abs_sin_le_one)
 done
 
 lemma abs_cos_le_one [simp]: "\<bar>cos x\<bar> \<le> 1"
 by (rule power2_le_imp_le, simp_all add: cos_squared_eq)
 
 lemma cos_ge_minus_one [simp]: "-1 \<le> cos x"
-apply (insert abs_cos_le_one [of x]) 
-apply (simp add: abs_le_iff del: abs_cos_le_one) 
+apply (insert abs_cos_le_one [of x])
+apply (simp add: abs_le_iff del: abs_cos_le_one)
 done
 
 lemma cos_le_one [simp]: "cos x \<le> 1"
-apply (insert abs_cos_le_one [of x]) 
+apply (insert abs_cos_le_one [of x])
 apply (simp add: abs_le_iff del: abs_cos_le_one)
 done
 
-lemma DERIV_fun_pow: "DERIV g x :> m ==>  
+lemma DERIV_fun_pow: "DERIV g x :> m ==>
       DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m"
 unfolding One_nat_def
 apply (rule lemma_DERIV_subst)
@@ -1515,15 +1516,15 @@
 
 (* lemma *)
 lemma lemma_DERIV_sin_cos_add:
-     "\<forall>x.  
-         DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
+     "\<forall>x.
+         DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +
                (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0"
   by (auto intro!: DERIV_intros simp add: algebra_simps)
 
 lemma sin_cos_add [simp]:
-     "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
+     "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +
       (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0"
-apply (cut_tac y = 0 and x = x and y7 = y 
+apply (cut_tac y = 0 and x = x and y7 = y
        in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all])
 apply (auto simp add: numeral_2_eq_2)
 done
@@ -1543,9 +1544,9 @@
   by (auto intro!: DERIV_intros simp add: algebra_simps)
 
 
-lemma sin_cos_minus: 
+lemma sin_cos_minus:
     "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0"
-apply (cut_tac y = 0 and x = x 
+apply (cut_tac y = 0 and x = x
        in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all])
 apply simp
 done
@@ -1582,27 +1583,27 @@
   pi :: "real" where
   "pi = 2 * (THE x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)"
 
-text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; 
+text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"};
    hence define pi.*}
 
 lemma sin_paired:
-     "(%n. -1 ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1)) 
+     "(%n. -1 ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1))
       sums  sin x"
 proof -
   have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. sin_coeff k * x ^ k) sums sin x"
     unfolding sin_def
-    by (rule sin_converges [THEN sums_summable, THEN sums_group], simp) 
+    by (rule sin_converges [THEN sums_summable, THEN sums_group], simp)
   thus ?thesis unfolding One_nat_def sin_coeff_def by (simp add: mult_ac)
 qed
 
 text {* FIXME: This is a long, ugly proof! *}
 lemma sin_gt_zero: "[|0 < x; x < 2 |] ==> 0 < sin x"
-apply (subgoal_tac 
+apply (subgoal_tac
        "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2.
-              -1 ^ k / real (fact (2 * k + 1)) * x ^ (2 * k + 1)) 
+              -1 ^ k / real (fact (2 * k + 1)) * x ^ (2 * k + 1))
      sums (\<Sum>n. -1 ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))")
  prefer 2
- apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) 
+ apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp)
 apply (rotate_tac 2)
 apply (drule sin_paired [THEN sums_unique, THEN ssubst])
 unfolding One_nat_def
@@ -1614,10 +1615,10 @@
 apply (erule sums_summable)
 apply (case_tac "m=0")
 apply (simp (no_asm_simp))
-apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") 
-apply (simp only: mult_less_cancel_left, simp)  
+apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x")
+apply (simp only: mult_less_cancel_left, simp)
 apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric])
-apply (subgoal_tac "x*x < 2*3", simp) 
+apply (subgoal_tac "x*x < 2*3", simp)
 apply (rule mult_strict_mono)
 apply (auto simp add: real_0_less_add_iff real_of_nat_Suc simp del: fact_Suc)
 apply (subst fact_Suc)
@@ -1630,15 +1631,15 @@
 apply (subst real_of_nat_mult)
 apply (simp (no_asm) add: divide_inverse del: fact_Suc)
 apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc)
-apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) 
+apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right)
 apply (auto simp add: mult_assoc simp del: fact_Suc)
-apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) 
+apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right)
 apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc)
-apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") 
+apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)")
 apply (erule ssubst)+
 apply (auto simp del: fact_Suc)
 apply (subgoal_tac "0 < x ^ (4 * m) ")
- prefer 2 apply (simp only: zero_less_power) 
+ prefer 2 apply (simp only: zero_less_power)
 apply (simp (no_asm_simp) add: mult_less_cancel_left)
 apply (rule mult_strict_mono)
 apply (simp_all (no_asm_simp))
@@ -1657,7 +1658,7 @@
 proof -
   have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. cos_coeff k * x ^ k) sums cos x"
     unfolding cos_def
-    by (rule cos_converges [THEN sums_summable, THEN sums_group], simp) 
+    by (rule cos_converges [THEN sums_summable, THEN sums_group], simp)
   thus ?thesis unfolding cos_coeff_def by (simp add: mult_ac)
 qed
 
@@ -1665,12 +1666,12 @@
 by simp
 
 lemma real_mult_inverse_cancel:
-     "[|(0::real) < x; 0 < x1; x1 * y < x * u |] 
+     "[|(0::real) < x; 0 < x1; x1 * y < x * u |]
       ==> inverse x * y < inverse x1 * u"
-apply (rule_tac c=x in mult_less_imp_less_left) 
+apply (rule_tac c=x in mult_less_imp_less_left)
 apply (auto simp add: mult_assoc [symmetric])
 apply (simp (no_asm) add: mult_ac)
-apply (rule_tac c=x1 in mult_less_imp_less_right) 
+apply (rule_tac c=x1 in mult_less_imp_less_right)
 apply (auto simp add: mult_ac)
 done
 
@@ -1687,7 +1688,7 @@
 lemma cos_two_less_zero [simp]: "cos (2) < 0"
 apply (cut_tac x = 2 in cos_paired)
 apply (drule sums_minus)
-apply (rule neg_less_iff_less [THEN iffD1]) 
+apply (rule neg_less_iff_less [THEN iffD1])
 apply (frule sums_unique, auto)
 apply (rule_tac y =
  "\<Sum>n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))"
@@ -1697,12 +1698,12 @@
 apply (rule sumr_pos_lt_pair)
 apply (erule sums_summable, safe)
 unfolding One_nat_def
-apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] 
+apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric]
             del: fact_Suc)
 apply (rule real_mult_inverse_cancel2)
 apply (rule real_of_nat_fact_gt_zero)+
 apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc)
-apply (subst fact_lemma) 
+apply (subst fact_lemma)
 apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"])
 apply (simp only: real_of_nat_mult)
 apply (rule mult_strict_mono, force)
@@ -1822,7 +1823,7 @@
 lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n"
 proof -
   have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute)
-  also have "... = -1 ^ n" by (rule cos_npi) 
+  also have "... = -1 ^ n" by (rule cos_npi)
   finally show ?thesis .
 qed
 
@@ -1832,7 +1833,7 @@
 done
 
 lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0"
-by (simp add: mult_commute [of pi]) 
+by (simp add: mult_commute [of pi])
 
 lemma cos_two_pi [simp]: "cos (2 * pi) = 1"
 by (simp add: cos_double)
@@ -1846,10 +1847,10 @@
 apply (rule pi_half_less_two)
 done
 
-lemma sin_less_zero: 
+lemma sin_less_zero:
   assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0"
 proof -
-  have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2) 
+  have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2)
   thus ?thesis by simp
 qed
 
@@ -1862,7 +1863,7 @@
 apply (cut_tac cos_is_zero, safe)
 apply (rename_tac y z)
 apply (drule_tac x = y in spec)
-apply (drule_tac x = "pi/2" in spec, simp) 
+apply (drule_tac x = "pi/2" in spec, simp)
 done
 
 lemma cos_gt_zero_pi: "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x"
@@ -1871,10 +1872,10 @@
 apply (rule cos_gt_zero)
 apply (auto intro: cos_gt_zero)
 done
- 
+
 lemma cos_ge_zero: "[| -(pi/2) \<le> x; x \<le> pi/2 |] ==> 0 \<le> cos x"
 apply (auto simp add: order_le_less cos_gt_zero_pi)
-apply (subgoal_tac "x = pi/2", auto) 
+apply (subgoal_tac "x = pi/2", auto)
 done
 
 lemma sin_gt_zero_pi: "[| 0 < x; x < pi  |] ==> 0 < sin x"
@@ -1897,7 +1898,7 @@
   qed
   then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast
   hence "0 < sin y" using sin_gt_zero by auto
-  moreover 
+  moreover
   have "sin y < 0" using sin_gt_zero_pi[of "y - pi"] `pi < y` and `y < 2 * pi` sin_periodic_pi[of "y - pi"] by auto
   ultimately show False by auto
 qed
@@ -1914,7 +1915,7 @@
 apply (drule_tac f = cos in Rolle)
 apply (drule_tac [5] f = cos in Rolle)
 apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos
-            dest!: DERIV_cos [THEN DERIV_unique] 
+            dest!: DERIV_cos [THEN DERIV_unique]
             simp add: differentiable_def)
 apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans])
 done
@@ -1925,32 +1926,32 @@
 apply (subgoal_tac "\<forall>x. (- (pi/2) \<le> x & x \<le> pi/2 & (sin x = y)) = (0 \<le> (x + pi/2) & (x + pi/2) \<le> pi & (cos (x + pi/2) = -y))")
 apply (erule contrapos_np)
 apply (simp del: minus_sin_cos_eq [symmetric])
-apply (cut_tac y="-y" in cos_total, simp) apply simp 
+apply (cut_tac y="-y" in cos_total, simp) apply simp
 apply (erule ex1E)
 apply (rule_tac a = "x - (pi/2)" in ex1I)
 apply (simp (no_asm) add: add_assoc)
 apply (rotate_tac 3)
-apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) 
+apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all)
 done
 
 lemma reals_Archimedean4:
      "[| 0 < y; 0 \<le> x |] ==> \<exists>n. real n * y \<le> x & x < real (Suc n) * y"
 apply (auto dest!: reals_Archimedean3)
-apply (drule_tac x = x in spec, clarify) 
+apply (drule_tac x = x in spec, clarify)
 apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y")
- prefer 2 apply (erule LeastI) 
-apply (case_tac "LEAST m::nat. x < real m * y", simp) 
+ prefer 2 apply (erule LeastI)
+apply (case_tac "LEAST m::nat. x < real m * y", simp)
 apply (subgoal_tac "~ x < real nat * y")
- prefer 2 apply (rule not_less_Least, simp, force)  
+ prefer 2 apply (rule not_less_Least, simp, force)
 done
 
-(* Pre Isabelle99-2 proof was simpler- numerals arithmetic 
+(* Pre Isabelle99-2 proof was simpler- numerals arithmetic
    now causes some unwanted re-arrangements of literals!   *)
 lemma cos_zero_lemma:
-     "[| 0 \<le> x; cos x = 0 |] ==>  
+     "[| 0 \<le> x; cos x = 0 |] ==>
       \<exists>n::nat. ~even n & x = real n * (pi/2)"
 apply (drule pi_gt_zero [THEN reals_Archimedean4], safe)
-apply (subgoal_tac "0 \<le> x - real n * pi & 
+apply (subgoal_tac "0 \<le> x - real n * pi &
                     (x - real n * pi) \<le> pi & (cos (x - real n * pi) = 0) ")
 apply (auto simp add: algebra_simps real_of_nat_Suc)
  prefer 2 apply (simp add: cos_diff)
@@ -1965,39 +1966,39 @@
 done
 
 lemma sin_zero_lemma:
-     "[| 0 \<le> x; sin x = 0 |] ==>  
+     "[| 0 \<le> x; sin x = 0 |] ==>
       \<exists>n::nat. even n & x = real n * (pi/2)"
 apply (subgoal_tac "\<exists>n::nat. ~ even n & x + pi/2 = real n * (pi/2) ")
  apply (clarify, rule_tac x = "n - 1" in exI)
  apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib)
 apply (rule cos_zero_lemma)
-apply (simp_all add: add_increasing)  
+apply (simp_all add: add_increasing)
 done
 
 
 lemma cos_zero_iff:
-     "(cos x = 0) =  
-      ((\<exists>n::nat. ~even n & (x = real n * (pi/2))) |    
+     "(cos x = 0) =
+      ((\<exists>n::nat. ~even n & (x = real n * (pi/2))) |
        (\<exists>n::nat. ~even n & (x = -(real n * (pi/2)))))"
 apply (rule iffI)
 apply (cut_tac linorder_linear [of 0 x], safe)
 apply (drule cos_zero_lemma, assumption+)
-apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) 
-apply (force simp add: minus_equation_iff [of x]) 
-apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) 
+apply (cut_tac x="-x" in cos_zero_lemma, simp, simp)
+apply (force simp add: minus_equation_iff [of x])
+apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib)
 apply (auto simp add: cos_add)
 done
 
 (* ditto: but to a lesser extent *)
 lemma sin_zero_iff:
-     "(sin x = 0) =  
-      ((\<exists>n::nat. even n & (x = real n * (pi/2))) |    
+     "(sin x = 0) =
+      ((\<exists>n::nat. even n & (x = real n * (pi/2))) |
        (\<exists>n::nat. even n & (x = -(real n * (pi/2)))))"
 apply (rule iffI)
 apply (cut_tac linorder_linear [of 0 x], safe)
 apply (drule sin_zero_lemma, assumption+)
 apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe)
-apply (force simp add: minus_equation_iff [of x]) 
+apply (force simp add: minus_equation_iff [of x])
 apply (auto simp add: even_mult_two_ex)
 done
 
@@ -2066,19 +2067,19 @@
 lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x"
 by (simp add: tan_def)
 
-lemma lemma_tan_add1: 
-      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]  
+lemma lemma_tan_add1:
+      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]
         ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)"
 apply (simp add: tan_def divide_inverse)
-apply (auto simp del: inverse_mult_distrib 
+apply (auto simp del: inverse_mult_distrib
             simp add: inverse_mult_distrib [symmetric] mult_ac)
 apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst])
-apply (auto simp del: inverse_mult_distrib 
+apply (auto simp del: inverse_mult_distrib
             simp add: mult_assoc left_diff_distrib cos_add)
 done
 
-lemma add_tan_eq: 
-      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]  
+lemma add_tan_eq:
+      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]
        ==> tan x + tan y = sin(x + y)/(cos x * cos y)"
 apply (simp add: tan_def)
 apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst])
@@ -2087,27 +2088,27 @@
 done
 
 lemma tan_add:
-     "[| cos x \<noteq> 0; cos y \<noteq> 0; cos (x + y) \<noteq> 0 |]  
+     "[| cos x \<noteq> 0; cos y \<noteq> 0; cos (x + y) \<noteq> 0 |]
       ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))"
 apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1)
 apply (simp add: tan_def)
 done
 
 lemma tan_double:
-     "[| cos x \<noteq> 0; cos (2 * x) \<noteq> 0 |]  
+     "[| cos x \<noteq> 0; cos (2 * x) \<noteq> 0 |]
       ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))"
-apply (insert tan_add [of x x]) 
-apply (simp add: mult_2 [symmetric])  
+apply (insert tan_add [of x x])
+apply (simp add: mult_2 [symmetric])
 apply (auto simp add: numeral_2_eq_2)
 done
 
 lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x"
-by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) 
-
-lemma tan_less_zero: 
+by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi)
+
+lemma tan_less_zero:
   assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0"
 proof -
-  have "0 < tan (- x)" using assms by (simp only: tan_gt_zero) 
+  have "0 < tan (- x)" using assms by (simp only: tan_gt_zero)
   thus ?thesis by simp
 qed
 
@@ -2143,8 +2144,8 @@
 apply (rule LIM_mult)
 apply (rule_tac [2] inverse_1 [THEN subst])
 apply (rule_tac [2] LIM_inverse)
-apply (simp_all add: divide_inverse [symmetric]) 
-apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) 
+apply (simp_all add: divide_inverse [symmetric])
+apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric])
 apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+
 done
 
@@ -2189,12 +2190,12 @@
 apply (subgoal_tac "\<exists>z. xa < z & z < y & DERIV tan z :> 0")
 apply (rule_tac [4] Rolle)
 apply (rule_tac [2] Rolle)
-apply (auto intro!: DERIV_tan DERIV_isCont exI 
+apply (auto intro!: DERIV_tan DERIV_isCont exI
             simp add: differentiable_def)
 txt{*Now, simulate TRYALL*}
 apply (rule_tac [!] DERIV_tan asm_rl)
 apply (auto dest!: DERIV_unique [OF _ DERIV_tan]
-            simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) 
+            simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym])
 done
 
 lemma tan_monotone: assumes "- (pi / 2) < y" and "y < x" and "x < pi / 2"
@@ -2208,7 +2209,7 @@
     have "cos x' \<noteq> 0" by auto
     thus "DERIV tan x' :> inverse (cos x'^2)" by (rule DERIV_tan)
   qed
-  from MVT2[OF `y < x` this] 
+  from MVT2[OF `y < x` this]
   obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<twosuperior>)" by auto
   hence "- (pi / 2) < z" and "z < pi / 2" using assms by auto
   hence "0 < cos z" using cos_gt_zero_pi by auto
@@ -2228,7 +2229,7 @@
   show "y < x"
   proof (rule ccontr)
     assume "\<not> y < x" hence "x \<le> y" by auto
-    hence "tan x \<le> tan y" 
+    hence "tan x \<le> tan y"
     proof (cases "x = y")
       case True thus ?thesis by auto
     next
@@ -2241,10 +2242,10 @@
 
 lemma tan_inverse: "1 / (tan y) = tan (pi / 2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto
 
-lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" 
+lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x"
   by (simp add: tan_def)
 
-lemma tan_periodic_nat[simp]: fixes n :: nat shows "tan (x + real n * pi) = tan x" 
+lemma tan_periodic_nat[simp]: fixes n :: nat shows "tan (x + real n * pi) = tan x"
 proof (induct n arbitrary: x)
   case (Suc n)
   have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 real_of_nat_add real_of_one left_distrib by auto
@@ -2275,18 +2276,18 @@
   arccos :: "real => real" where
   "arccos y = (THE x. 0 \<le> x & x \<le> pi & cos x = y)"
 
-definition     
+definition
   arctan :: "real => real" where
   "arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)"
 
 lemma arcsin:
-     "[| -1 \<le> y; y \<le> 1 |]  
-      ==> -(pi/2) \<le> arcsin y &  
+     "[| -1 \<le> y; y \<le> 1 |]
+      ==> -(pi/2) \<le> arcsin y &
            arcsin y \<le> pi/2 & sin(arcsin y) = y"
 unfolding arcsin_def by (rule theI' [OF sin_total])
 
 lemma arcsin_pi:
-     "[| -1 \<le> y; y \<le> 1 |]  
+     "[| -1 \<le> y; y \<le> 1 |]
       ==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi & sin(arcsin y) = y"
 apply (drule (1) arcsin)
 apply (force intro: order_trans)
@@ -2294,7 +2295,7 @@
 
 lemma sin_arcsin [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> sin(arcsin y) = y"
 by (blast dest: arcsin)
-      
+
 lemma arcsin_bounded:
      "[| -1 \<le> y; y \<le> 1 |] ==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi/2"
 by (blast dest: arcsin)
@@ -2323,13 +2324,13 @@
 done
 
 lemma arccos:
-     "[| -1 \<le> y; y \<le> 1 |]  
+     "[| -1 \<le> y; y \<le> 1 |]
       ==> 0 \<le> arccos y & arccos y \<le> pi & cos(arccos y) = y"
 unfolding arccos_def by (rule theI' [OF cos_total])
 
 lemma cos_arccos [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> cos(arccos y) = y"
 by (blast dest: arccos)
-      
+
 lemma arccos_bounded: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y & arccos y \<le> pi"
 by (blast dest: arccos)
 
@@ -2340,7 +2341,7 @@
 by (blast dest: arccos)
 
 lemma arccos_lt_bounded:
-     "[| -1 < y; y < 1 |]  
+     "[| -1 < y; y < 1 |]
       ==> 0 < arccos y & arccos y < pi"
 apply (frule order_less_imp_le)
 apply (frule_tac y = y in order_less_imp_le)
@@ -2400,7 +2401,7 @@
 lemma arctan_ubound: "arctan y < pi/2"
 by (auto simp only: arctan)
 
-lemma arctan_tan: 
+lemma arctan_tan:
       "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x"
 apply (unfold arctan_def)
 apply (rule the1_equality)
@@ -2415,7 +2416,7 @@
 apply (case_tac "n")
 apply (case_tac [3] "n")
 apply (cut_tac [2] y = x in arctan_ubound)
-apply (cut_tac [4] y = x in arctan_lbound) 
+apply (cut_tac [4] y = x in arctan_lbound)
 apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff)
 done
 
@@ -2423,7 +2424,7 @@
 apply (rule power_inverse [THEN subst])
 apply (rule_tac c1 = "(cos x)\<twosuperior>" in real_mult_right_cancel [THEN iffD1])
 apply (auto dest: field_power_not_zero
-        simp add: power_mult_distrib left_distrib power_divide tan_def 
+        simp add: power_mult_distrib left_distrib power_divide tan_def
                   mult_assoc power_inverse [symmetric])
 done
 
@@ -2588,7 +2589,7 @@
   have "sin ((real n + 1/2) * pi) = cos (real n * pi)"
     by (auto simp add: algebra_simps sin_add)
   thus ?thesis
-    by (simp add: real_of_nat_Suc left_distrib add_divide_distrib 
+    by (simp add: real_of_nat_Suc left_distrib add_divide_distrib
                   mult_commute [of pi])
 qed
 
@@ -2627,7 +2628,7 @@
 proof -
   obtain z where "- (pi / 2) < z" and "z < pi / 2" and "tan z = x" using tan_total by blast
   have "tan (pi / 4) = 1" and "tan (- (pi / 4)) = - 1" using tan_45 tan_minus by auto
-  have "z \<noteq> pi / 4" 
+  have "z \<noteq> pi / 4"
   proof (rule ccontr)
     assume "\<not> (z \<noteq> pi / 4)" hence "z = pi / 4" by auto
     have "tan z = 1" unfolding `z = pi / 4` `tan (pi / 4) = 1` ..
@@ -2644,11 +2645,11 @@
   proof (rule ccontr)
     assume "\<not> (z < pi / 4)" hence "pi / 4 < z" using `z \<noteq> pi / 4` by auto
     have "- (pi / 2) < pi / 4" using m2pi_less_pi by auto
-    from tan_monotone[OF this `pi / 4 < z` `z < pi / 2`] 
+    from tan_monotone[OF this `pi / 4 < z` `z < pi / 2`]
     have "1 < x" unfolding `tan z = x` `tan (pi / 4) = 1` .
     thus False using `\<bar>x\<bar> < 1` by auto
   qed
-  moreover 
+  moreover
   have "-(pi / 4) < z"
   proof (rule ccontr)
     assume "\<not> (-(pi / 4) < z)" hence "z < - (pi / 4)" using `z \<noteq> - (pi / 4)` by auto
@@ -2677,14 +2678,14 @@
     show ?thesis
     proof (cases "x = 1")
       case True hence "tan (pi/4) = x" using tan_45 by auto
-      moreover 
+      moreover
       have "- pi \<le> pi" unfolding minus_le_self_iff by auto
       hence "-(pi/4) \<le> pi/4" and "pi/4 \<le> pi/4" by auto
       ultimately show ?thesis by blast
     next
       case False hence "x = -1" using `\<not> \<bar>x\<bar> < 1` and `\<bar>x\<bar> \<le> 1` by auto
       hence "tan (-(pi/4)) = x" using tan_45 tan_minus by auto
-      moreover 
+      moreover
       have "- pi \<le> pi" unfolding minus_le_self_iff by auto
       hence "-(pi/4) \<le> pi/4" and "-(pi/4) \<le> -(pi/4)" by auto
       ultimately show ?thesis by blast
@@ -2723,7 +2724,7 @@
   have "\<bar>5 / 12\<bar> < (1 :: real)" by auto
   from arctan_add[OF less_imp_le[OF this] this]
   have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto
-  moreover 
+  moreover
   have "\<bar>1\<bar> \<le> (1::real)" and "\<bar>1 / 239\<bar> < (1::real)" by auto
   from arctan_add[OF this]
   have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" by auto
@@ -2749,7 +2750,7 @@
         show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
       qed
     } note mono = this
-    
+
     show ?thesis
     proof (cases "0 \<le> x")
       case True from mono[OF this `x \<le> 1`, THEN allI]
@@ -2793,7 +2794,7 @@
   from mult_left_mono[OF less_imp_le[OF `\<bar>x\<bar> < 1`] abs_ge_zero[of x]]
   have "\<bar> x^2 \<bar> < 1" using `\<bar> x \<bar> < 1` unfolding numeral_2_eq_2 power_Suc2 by auto
   thus ?thesis using zero_le_power2 by auto
-qed 
+qed
 
 lemma DERIV_arctan_series: assumes "\<bar> x \<bar> < 1"
   shows "DERIV (\<lambda> x'. \<Sum> k. (-1)^k * (1 / real (k*2+1) * x' ^ (k*2+1))) x :> (\<Sum> k. (-1)^k * x^(k*2))" (is "DERIV ?arctan _ :> ?Int")
@@ -2812,7 +2813,7 @@
   { fix f :: "nat \<Rightarrow> real"
     have "\<And> x. f sums x = (\<lambda> n. if even n then f (n div 2) else 0) sums x"
     proof
-      fix x :: real assume "f sums x" 
+      fix x :: real assume "f sums x"
       from sums_if[OF sums_zero this]
       show "(\<lambda> n. if even n then f (n div 2) else 0) sums x" by auto
     next
@@ -2827,10 +2828,10 @@
     by auto
 
   { fix x :: real
-    have if_eq': "\<And> n. (if even n then -1 ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = 
+    have if_eq': "\<And> n. (if even n then -1 ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n =
       (if even n then -1 ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)"
       using n_even by auto
-    have idx_eq: "\<And> n. n * 2 + 1 = Suc (2 * n)" by auto 
+    have idx_eq: "\<And> n. n * 2 + 1 = Suc (2 * n)" by auto
     have "(\<Sum> n. ?f n * x^(Suc n)) = ?arctan x" unfolding if_eq' idx_eq suminf_def sums_even[of "\<lambda> n. -1 ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric]
       by auto
   } note arctan_eq = this
@@ -2893,10 +2894,10 @@
         show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y" using DERIV_in_rball DERIV_isCont by auto
       qed
     qed
-    
+
     have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0"
       unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto
-    
+
     have "suminf (?c x) - arctan x = 0"
     proof (cases "x = 0")
       case True thus ?thesis using suminf_arctan_zero by auto
@@ -2909,7 +2910,7 @@
       have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
         by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"])
           (simp_all only: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>` neg_less_iff_less)
-      ultimately 
+      ultimately
       show ?thesis using suminf_arctan_zero by auto
     qed
     thus ?thesis by auto
@@ -2971,16 +2972,16 @@
     from this[unfolded LIMSEQ_rabs_zero diff_minus add_commute[of "arctan 1"], THEN LIMSEQ_add_const, of "- arctan 1", THEN LIMSEQ_minus]
     have "(?c 1) sums (arctan 1)" unfolding sums_def by auto
     hence "arctan 1 = (\<Sum> i. ?c 1 i)" by (rule sums_unique)
-    
+
     show ?thesis
     proof (cases "x = 1", simp add: `arctan 1 = (\<Sum> i. ?c 1 i)`)
       assume "x \<noteq> 1" hence "x = -1" using `\<bar>x\<bar> = 1` by auto
-      
+
       have "- (pi / 2) < 0" using pi_gt_zero by auto
       have "- (2 * pi) < 0" using pi_gt_zero by auto
-      
+
       have c_minus_minus: "\<And> i. ?c (- 1) i = - ?c 1 i" unfolding One_nat_def by auto
-    
+
       have "arctan (- 1) = arctan (tan (-(pi / 4)))" unfolding tan_45 tan_minus ..
       also have "\<dots> = - (pi / 4)" by (rule arctan_tan, auto simp add: order_less_trans[OF `- (pi / 2) < 0` pi_gt_zero])
       also have "\<dots> = - (arctan (tan (pi / 4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric], auto simp add: order_less_trans[OF `- (2 * pi) < 0` pi_gt_zero])
@@ -2999,7 +3000,7 @@
   hence low2: "- (pi / 2) < y / 2" and high2: "y / 2 < pi / 2" by auto
 
   have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> A / B = (A / C) / (B / C)" by auto
-  
+
   have "0 < cos y" using cos_gt_zero_pi[OF low high] .
   hence "cos y \<noteq> 0" and cos_sqrt: "sqrt ((cos y) ^ 2) = cos y" by auto
 
@@ -3032,14 +3033,14 @@
 qed
 
 lemma arctan_monotone': assumes "x \<le> y" shows "arctan x \<le> arctan y"
-proof (cases "x = y") 
+proof (cases "x = y")
   case False hence "x < y" using `x \<le> y` by auto from arctan_monotone[OF this] show ?thesis by auto
 qed auto
 
-lemma arctan_minus: "arctan (- x) = - arctan x" 
+lemma arctan_minus: "arctan (- x) = - arctan x"
 proof -
   obtain y where "- (pi / 2) < y" and "y < pi / 2" and "tan y = x" using tan_total by blast
-  thus ?thesis unfolding `tan y = x`[symmetric] tan_minus[symmetric] using arctan_tan[of y] arctan_tan[of "-y"] by auto 
+  thus ?thesis unfolding `tan y = x`[symmetric] tan_minus[symmetric] using arctan_tan[of y] arctan_tan[of "-y"] by auto
 qed
 
 lemma arctan_inverse: assumes "x \<noteq> 0" shows "arctan (1 / x) = sgn x * pi / 2 - arctan x"
@@ -3062,7 +3063,7 @@
     case True from pos_y[OF this `y < pi / 2` `y = arctan x` `tan y = x`] show ?thesis .
   next
     case False hence "y \<le> 0" by auto
-    moreover have "y \<noteq> 0" 
+    moreover have "y \<noteq> 0"
     proof (rule ccontr)
       assume "\<not> y \<noteq> 0" hence "y = 0" by auto
       have "x = 0" unfolding `tan y = x`[symmetric] `y = 0` tan_zero ..