it is known as the extended reals, not the infinite reals
authorhoelzl
Fri, 03 Dec 2010 15:25:14 +0100
changeset 41023 9118eb4eb8dc
parent 41022 81d337539d57
child 41024 ba961a606c67
it is known as the extended reals, not the infinite reals
NEWS
src/HOL/IsaMakefile
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Complete_Measure.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Lebesgue_Integration.thy
src/HOL/Probability/Lebesgue_Measure.thy
src/HOL/Probability/Measure.thy
src/HOL/Probability/Positive_Extended_Real.thy
src/HOL/Probability/Positive_Infinite_Real.thy
src/HOL/Probability/Probability_Space.thy
src/HOL/Probability/Product_Measure.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/ex/Dining_Cryptographers.thy
src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy
--- a/NEWS	Mon Dec 06 19:18:02 2010 +0100
+++ b/NEWS	Fri Dec 03 15:25:14 2010 +0100
@@ -334,8 +334,8 @@
 of euclidean spaces the real and complex numbers are instantiated to
 be euclidean_spaces.  INCOMPATIBILITY.
 
-* Probability: Introduced pinfreal as real numbers with infinity. Use
-pinfreal as value for measures. Introduce the Radon-Nikodym
+* Probability: Introduced pextreal as positive extended real numbers.
+Use pextreal as value for measures. Introduce the Radon-Nikodym
 derivative, product spaces and Fubini's theorem for arbitrary sigma
 finite measures. Introduces Lebesgue measure based on the integral in
 Multivariate Analysis.  INCOMPATIBILITY.
--- a/src/HOL/IsaMakefile	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/IsaMakefile	Fri Dec 03 15:25:14 2010 +0100
@@ -1183,7 +1183,7 @@
   Probability/ex/Koepf_Duermuth_Countermeasure.thy			\
   Probability/Information.thy Probability/Lebesgue_Integration.thy	\
   Probability/Lebesgue_Measure.thy Probability/Measure.thy		\
-  Probability/Positive_Infinite_Real.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			\
--- a/src/HOL/Probability/Borel_Space.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Borel_Space.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -3,7 +3,7 @@
 header {*Borel spaces*}
 
 theory Borel_Space
-  imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
+  imports Sigma_Algebra Positive_Extended_Real Multivariate_Analysis
 begin
 
 lemma LIMSEQ_max:
@@ -1012,10 +1012,10 @@
 lemma borel_Real_measurable:
   "A \<in> sets borel \<Longrightarrow> Real -` A \<in> sets borel"
 proof (rule borel_measurable_translate)
-  fix B :: "pinfreal set" assume "open B"
+  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_pinfreal_def by blast
+    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..})"
@@ -1027,7 +1027,7 @@
 qed simp
 
 lemma borel_real_measurable:
-  "A \<in> sets borel \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel"
+  "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 }
@@ -1035,10 +1035,10 @@
   have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
     by (force simp: Ex_less_real)
 
-  have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
-    unfolding open_pinfreal_def using `open B`
+  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 :: pinfreal set) \<in> sets borel" unfolding * by auto
+  then show "(real -` B :: pextreal set) \<in> sets borel" unfolding * by auto
 qed simp
 
 lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
@@ -1046,7 +1046,7 @@
   shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
   unfolding in_borel_measurable_borel
 proof safe
-  fix S :: "pinfreal set" assume "S \<in> sets borel"
+  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
@@ -1056,7 +1056,7 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+  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
@@ -1085,7 +1085,7 @@
     by (simp cong: measurable_cong)
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real:
+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
@@ -1130,8 +1130,8 @@
   ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
 qed
 
-lemma (in sigma_algebra) less_eq_le_pinfreal_measurable:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+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)"
 proof
   assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
@@ -1143,9 +1143,9 @@
       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_pinfreal_inverse_of_nat_Suc_less[of "f x - a"]
+        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: pinfreal_minus_order)
+          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
@@ -1174,7 +1174,7 @@
         have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
         proof safe
           fix x assume "f x < a" and [simp]: "x \<in> space M"
-          with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"]
+          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)) "
@@ -1197,7 +1197,7 @@
         show "f x = \<omega>"    proof (rule ccontr)
           assume "f x \<noteq> \<omega>"
           with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
-            by (auto simp: pinfreal_noteq_omega_Ex)
+            by (auto simp: pextreal_noteq_omega_Ex)
           with *[THEN spec, of n] show False by auto
         qed
       qed
@@ -1209,8 +1209,8 @@
   qed
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
+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
@@ -1219,9 +1219,9 @@
 next
   assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
   hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
-    unfolding less_eq_le_pinfreal_measurable
+    unfolding less_eq_le_pextreal_measurable
     unfolding greater_eq_le_measurable .
-  show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater
+  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
@@ -1242,28 +1242,28 @@
   qed
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable .
+lemma (in sigma_algebra) borel_measurable_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_pinfreal_iff_le:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable .
+lemma (in sigma_algebra) borel_measurable_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_pinfreal_iff_ge:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable .
+lemma (in sigma_algebra) borel_measurable_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_pinfreal_eq_const:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_pextreal_eq_const:
+  fixes f :: "'a \<Rightarrow> pextreal" assumes "f \<in> borel_measurable M"
   shows "{x\<in>space M. f x = c} \<in> sets M"
 proof -
   have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
   then show ?thesis using assms by (auto intro!: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_neq_const:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_neq_const:
+  fixes f :: "'a \<Rightarrow> pextreal"
   assumes "f \<in> borel_measurable M"
   shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 proof -
@@ -1271,8 +1271,8 @@
   then show ?thesis using assms by (auto intro!: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_less[intro,simp]:
+  fixes f g :: "'a \<Rightarrow> pextreal"
   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"
@@ -1282,17 +1282,17 @@
     using assms by (auto intro!: borel_measurable_real)
   from borel_measurable_less[OF this]
   have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
-  moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
-  moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
-  moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
+  moreover have "{x \<in> space M. f x \<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_pinfreal_strict_mono_iff)
+    by (auto simp: real_of_pextreal_strict_mono_iff)
   ultimately show ?thesis by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+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"
@@ -1301,8 +1301,8 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_eq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pextreal"
   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"
@@ -1311,8 +1311,8 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_neq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pextreal"
   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"
@@ -1321,32 +1321,32 @@
   thus ?thesis using f g by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_add[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pextreal"
   assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
 proof -
   have *: "(\<lambda>x. f x + g x) =
      (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
-     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
+     by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex)
   show ?thesis using assms unfolding *
     by (auto intro!: measurable_If)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+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"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
 proof -
   have *: "(\<lambda>x. f x * g x) =
      (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
       Real (real (f x) * real (g x)))"
-     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
+     by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex)
   show ?thesis using assms unfolding *
     by (auto intro!: measurable_If)
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pextreal"
   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
@@ -1355,56 +1355,56 @@
     by induct auto
 qed (simp add: borel_measurable_const)
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_min[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_min[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> pextreal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
   using assms unfolding min_def by (auto intro!: measurable_If)
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_max[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_max[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> pextreal"
   assumes "f \<in> borel_measurable M"
   and "g \<in> borel_measurable M"
   shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
   using assms unfolding max_def by (auto intro!: measurable_If)
 
 lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
-  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pextreal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
-  unfolding borel_measurable_pinfreal_iff_greater
+  unfolding borel_measurable_pextreal_iff_greater
 proof safe
   fix a
   have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
-    by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'c=pinfreal])
+    by (auto simp: less_Sup_iff SUPR_def[where 'a=pextreal] SUPR_fun_expand[where 'c=pextreal])
   then show "{x\<in>space M. a < ?sup x} \<in> sets M"
     using assms by auto
 qed
 
 lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
-  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pextreal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
-  unfolding borel_measurable_pinfreal_iff_less
+  unfolding borel_measurable_pextreal_iff_less
 proof safe
   fix a
   have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
-    by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand)
+    by (auto simp: Inf_less_iff INFI_def[where 'a=pextreal] INFI_fun_expand)
   then show "{x\<in>space M. ?inf x < a} \<in> sets M"
     using assms by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_pinfreal_diff[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
+lemma (in sigma_algebra) borel_measurable_pextreal_diff[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> pextreal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  unfolding borel_measurable_pinfreal_iff_greater
+  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: pinfreal_less_minus_iff)
+    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
--- a/src/HOL/Probability/Caratheodory.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Caratheodory.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -1,14 +1,14 @@
 header {*Caratheodory Extension Theorem*}
 
 theory Caratheodory
-  imports Sigma_Algebra Positive_Infinite_Real
+  imports Sigma_Algebra Positive_Extended_Real
 begin
 
 text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
 
 subsection {* Measure Spaces *}
 
-definition "positive f \<longleftrightarrow> f {} = (0::pinfreal)" -- "Positive is enforced by the type"
+definition "positive f \<longleftrightarrow> f {} = (0::pextreal)" -- "Positive is enforced by the type"
 
 definition
   additive  where
@@ -58,7 +58,7 @@
      {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}"
 
 locale measure_space = sigma_algebra +
-  fixes \<mu> :: "'a set \<Rightarrow> pinfreal"
+  fixes \<mu> :: "'a set \<Rightarrow> pextreal"
   assumes empty_measure [simp]: "\<mu> {} = 0"
       and ca: "countably_additive M \<mu>"
 
@@ -148,7 +148,7 @@
   by (simp add:  lambda_system_def)
 
 lemma (in algebra) lambda_system_Compl:
-  fixes f:: "'a set \<Rightarrow> pinfreal"
+  fixes f:: "'a set \<Rightarrow> pextreal"
   assumes x: "x \<in> lambda_system M f"
   shows "space M - x \<in> lambda_system M f"
   proof -
@@ -161,7 +161,7 @@
   qed
 
 lemma (in algebra) lambda_system_Int:
-  fixes f:: "'a set \<Rightarrow> pinfreal"
+  fixes f:: "'a set \<Rightarrow> pextreal"
   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 -
@@ -196,7 +196,7 @@
 
 
 lemma (in algebra) lambda_system_Un:
-  fixes f:: "'a set \<Rightarrow> pinfreal"
+  fixes f:: "'a set \<Rightarrow> pextreal"
   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 -
@@ -295,7 +295,7 @@
   by (auto simp add: countably_subadditive_def o_def)
 
 lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> pinfreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> pextreal"
   assumes f: "positive f" and ad: "additive M f"
       and inc: "increasing M f"
       and A: "range A \<subseteq> sets M"
@@ -315,7 +315,7 @@
   by (simp add: increasing_def lambda_system_def)
 
 lemma (in algebra) lambda_system_strong_sum:
-  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pinfreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pextreal"
   assumes f: "positive f" and a: "a \<in> sets M"
       and A: "range A \<subseteq> lambda_system M f"
       and disj: "disjoint_family A"
@@ -497,7 +497,7 @@
   assumes posf: "positive f" and ca: "countably_additive M f"
       and s: "s \<in> sets M"
   shows "Inf (measure_set M f s) = f s"
-  unfolding Inf_pinfreal_def
+  unfolding Inf_pextreal_def
 proof (safe intro!: Greatest_equality)
   fix z
   assume z: "z \<in> measure_set M f s"
@@ -608,8 +608,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 pinfreal_le_epsilon)
-  fix A :: "nat \<Rightarrow> 'a set" and e :: pinfreal
+proof (safe, simp, rule pextreal_le_epsilon)
+  fix A :: "nat \<Rightarrow> 'a set" and e :: pextreal
 
   let "?outer n" = "Inf (measure_set M f (A n))"
   assume A: "range A \<subseteq> Pow (space M)"
@@ -688,8 +688,8 @@
     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 pinfreal_le_epsilon)
-      fix e :: pinfreal
+    proof (rule pextreal_le_epsilon)
+      fix e :: pextreal
       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"
@@ -760,7 +760,7 @@
 
 theorem (in algebra) caratheodory:
   assumes posf: "positive f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pinfreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma M) \<mu>"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pextreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma M) \<mu>"
   proof -
     have inc: "increasing M f"
       by (metis additive_increasing ca countably_additive_additive posf)
--- a/src/HOL/Probability/Complete_Measure.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -243,7 +243,7 @@
 qed
 
 lemma (in completeable_measure_space) completion_ex_borel_measurable:
-  fixes g :: "'a \<Rightarrow> pinfreal"
+  fixes g :: "'a \<Rightarrow> pextreal"
   assumes g: "g \<in> borel_measurable completion"
   shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
 proof -
--- a/src/HOL/Probability/Information.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Information.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -210,7 +210,7 @@
   have ms: "measure_space M \<nu>" by fact
   show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum"
     using RN_deriv_finite_measure[OF ms ac]
-    by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric])
+    by (auto intro!: setsum_cong simp: field_simps real_of_pextreal_mult[symmetric])
 qed
 
 lemma (in finite_prob_space) KL_divergence_positive_finite:
@@ -285,7 +285,7 @@
   note jd_commute = this
 
   { fix A assume A: "A \<in> sets (sigma (pair_algebra T S))"
-    have *: "\<And>x y. indicator ((\<lambda>(x, y). (y, x)) ` A) (x, y) = (indicator A (y, x) :: pinfreal)"
+    have *: "\<And>x y. indicator ((\<lambda>(x, y). (y, x)) ` A) (x, y) = (indicator A (y, x) :: pextreal)"
       unfolding indicator_def by auto
     have "ST.pair_measure ((\<lambda>(x, y). (y, x)) ` A) = TS.pair_measure A"
       unfolding ST.pair_measure_def TS.pair_measure_def
@@ -361,7 +361,7 @@
   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: pair_algebra_def setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
+       (auto simp add: pair_algebra_def setsum_cartesian_product' real_of_pextreal_mult[symmetric])
 
   show ?positive
     using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
@@ -463,7 +463,7 @@
     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 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_pinfreal_eq_0)
+    by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) (auto simp: real_of_pextreal_eq_0)
   finally show ?thesis
     using `simple_function X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
 qed
@@ -610,14 +610,14 @@
     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_pinfreal_mult[symmetric] real_of_pinfreal_eq_0)
+      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_pinfreal
-                       real_of_pinfreal_eq_0 zero_less_mult_iff)
+           (auto simp: zero_less_divide_iff zero_less_real_of_pextreal
+                       real_of_pextreal_eq_0 zero_less_mult_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))) -
@@ -721,7 +721,7 @@
   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_pinfreal_mult[symmetric])
+    by (intro setsum_cong) (auto intro!: arg_cong[where f="log b"] simp: real_of_pextreal_mult[symmetric])
   also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
     unfolding split_beta
   proof (rule log_setsum_divide)
@@ -743,15 +743,15 @@
     fix x assume "x \<in> ?M"
     let ?x = "(fst x, fst (snd x), snd (snd x))"
 
-    show "0 \<le> ?dXYZ {?x}" using real_pinfreal_nonneg .
+    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_pinfreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg)
+     by (simp add: real_pextreal_nonneg 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_pinfreal zero_less_mult_iff zero_less_divide_iff)
+         (auto simp add: zero_less_real_of_pextreal zero_less_mult_iff zero_less_divide_iff)
   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')
@@ -817,11 +817,11 @@
     also have "\<dots> = real (?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_pinfreal_mult[symmetric])
+      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]
       using finite_distribution_order(6)[OF MX MZ]
-      by (auto simp: log_simps field_simps zero_less_mult_iff zero_less_real_of_pinfreal real_of_pinfreal_eq_0)
+      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))" . }
   note * = this
@@ -830,7 +830,7 @@
     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_pinfreal_def
+                   setsum_commute[of _ "space MZ"] *   simp del: divide_pextreal_def
              intro!: setsum_cong)
 qed
 
@@ -853,7 +853,7 @@
   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_pinfreal_eq_0
+  by (auto simp: setsum_cartesian_product'  setsum_commute[of _ "Y`space M"] setsum_right_distrib real_of_pextreal_eq_0
            intro!: setsum_cong)
 
 lemma (in information_space) conditional_entropy_eq_cartesian_product:
@@ -880,8 +880,8 @@
   { 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_pinfreal_mult[symmetric] zero_less_mult_iff
-                     zero_less_real_of_pinfreal field_simps real_of_pinfreal_eq_0 abs_mult) }
+      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) }
   note * = this
   show ?thesis
     unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z]
@@ -913,8 +913,8 @@
   { 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_pinfreal_mult[symmetric] zero_less_mult_iff
-                     zero_less_real_of_pinfreal field_simps real_of_pinfreal_eq_0 abs_mult) }
+      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) }
   note * = this
   show ?thesis
     using setsum_real_joint_distribution_singleton[OF fY fX]
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -54,7 +54,7 @@
 qed
 
 lemma (in sigma_algebra) simple_function_indicator_representation:
-  fixes f ::"'a \<Rightarrow> pinfreal"
+  fixes f ::"'a \<Rightarrow> pextreal"
   assumes f: "simple_function 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")
@@ -69,7 +69,7 @@
 qed
 
 lemma (in measure_space) simple_function_notspace:
-  "simple_function (\<lambda>x. h x * indicator (- space M) x::pinfreal)" (is "simple_function ?h")
+  "simple_function (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function ?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)
@@ -212,7 +212,7 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
-  fixes u :: "'a \<Rightarrow> pinfreal"
+  fixes u :: "'a \<Rightarrow> pextreal"
   assumes u: "u \<in> borel_measurable M"
   shows "\<exists>f. (\<forall>i. simple_function (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
 proof -
@@ -265,7 +265,7 @@
     qed simp }
   note f_upper = this
 
-  let "?g j x" = "of_nat (f x j) / 2^j :: pinfreal"
+  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
@@ -350,7 +350,7 @@
     hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto
 
     show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
-    proof (rule pinfreal_SUPI)
+    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"
@@ -362,7 +362,7 @@
              (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff)
       qed
     next
-      fix y :: pinfreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
+      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)
@@ -404,7 +404,7 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
-  fixes u :: "'a \<Rightarrow> pinfreal"
+  fixes u :: "'a \<Rightarrow> pextreal"
   assumes "u \<in> borel_measurable M"
   obtains (x) f where "f \<up> u" "\<And>i. simple_function (f i)" "\<And>i. \<omega>\<notin>f i`space M"
 proof -
@@ -416,7 +416,7 @@
 qed
 
 lemma (in sigma_algebra) simple_function_eq_borel_measurable:
-  fixes f :: "'a \<Rightarrow> pinfreal"
+  fixes f :: "'a \<Rightarrow> pextreal"
   shows "simple_function f \<longleftrightarrow>
     finite (f`space M) \<and> f \<in> borel_measurable M"
   using simple_function_borel_measurable[of f]
@@ -424,7 +424,7 @@
   by (fastsimp simp: simple_function_def)
 
 lemma (in measure_space) simple_function_restricted:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
+  fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M"
   shows "sigma_algebra.simple_function (restricted_space A) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
     (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f")
 proof -
@@ -448,7 +448,7 @@
         using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
     next
       fix x
-      assume "indicator A x \<noteq> (0::pinfreal)"
+      assume "indicator A x \<noteq> (0::pextreal)"
       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
@@ -472,7 +472,7 @@
   by auto
 
 lemma (in sigma_algebra) simple_function_vimage:
-  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a"
   assumes g: "simple_function g" and f: "f \<in> S \<rightarrow> space M"
   shows "sigma_algebra.simple_function (vimage_algebra S f) (\<lambda>x. g (f x))"
 proof -
@@ -751,7 +751,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::pinfreal}" by auto
+  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pextreal}" by auto
   thus ?thesis
     using simple_integral_indicator[OF assms simple_function_const[of 1]]
     using sets_into_space[OF assms]
@@ -762,7 +762,7 @@
   assumes "simple_function u" "N \<in> null_sets"
   shows "simple_integral (\<lambda>x. u x * indicator N x) = 0"
 proof -
-  have "AE x. indicator N x = (0 :: pinfreal)"
+  have "AE x. indicator N x = (0 :: pextreal)"
     using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
   then have "simple_integral (\<lambda>x. u x * indicator N x) = simple_integral (\<lambda>x. 0)"
     using assms by (intro simple_integral_cong_AE) (auto intro!: AE_disjI2)
@@ -806,7 +806,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 pinfreal_mult_cancel_left by auto
+    unfolding pextreal_mult_cancel_left by auto
 qed
 
 lemma (in measure_space) simple_integral_subalgebra[simp]:
@@ -816,7 +816,7 @@
   unfolding measure_space.simple_integral_def_raw[OF assms] by simp
 
 lemma (in measure_space) simple_integral_vimage:
-  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a"
   assumes f: "bij_betw f S (space M)"
   shows "simple_integral g =
          measure_space.simple_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))"
@@ -893,7 +893,7 @@
             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: pinfreal_noteq_omega_Ex field_simps)
+            by (auto simp: pextreal_noteq_omega_Ex field_simps)
           also have "\<dots> = simple_integral ?g" using g `space M \<noteq> {}`
             by (subst simple_integral_indicator)
                (auto simp: image_constant ac_simps dest: simple_functionD)
@@ -950,7 +950,7 @@
   assumes "simple_function f"
   shows "positive_integral f = simple_integral f"
   unfolding positive_integral_def
-proof (safe intro!: pinfreal_SUPI)
+proof (safe intro!: pextreal_SUPI)
   fix g assume "simple_function g" "g \<le> f"
   with assms show "simple_integral g \<le> simple_integral f"
     by (auto intro!: simple_integral_mono simp: le_fun_def)
@@ -1017,7 +1017,7 @@
   using assms by blast
 
 lemma (in measure_space) positive_integral_vimage:
-  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a"
   assumes f: "bij_betw f S (space M)"
   shows "positive_integral g =
          measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))"
@@ -1039,14 +1039,14 @@
   show ?thesis
     unfolding positive_integral_alt1 T.positive_integral_alt1 SUPR_def * image_compose
   proof (safe intro!: arg_cong[where f=Sup] image_set_cong, simp_all add: comp_def)
-    fix g' :: "'a \<Rightarrow> pinfreal" assume "simple_function g'" "\<forall>x\<in>space M. g' x \<le> g x \<and> g' x \<noteq> \<omega>"
+    fix g' :: "'a \<Rightarrow> pextreal" assume "simple_function g'" "\<forall>x\<in>space M. g' x \<le> g x \<and> g' x \<noteq> \<omega>"
     then show "\<exists>h. T.simple_function h \<and> (\<forall>x\<in>S. h x \<le> g (f x) \<and> h x \<noteq> \<omega>) \<and>
                    T.simple_integral (\<lambda>x. g' (f x)) = T.simple_integral h"
       using f unfolding bij_betw_def
       by (auto intro!: exI[of _ "\<lambda>x. g' (f x)"]
                simp add: le_fun_def simple_function_vimage[OF _ f_fun])
   next
-    fix g' :: "'d \<Rightarrow> pinfreal" assume g': "T.simple_function g'" "\<forall>x\<in>S. g' x \<le> g (f x) \<and> g' x \<noteq> \<omega>"
+    fix g' :: "'d \<Rightarrow> pextreal" assume g': "T.simple_function g'" "\<forall>x\<in>S. g' x \<le> g (f x) \<and> g' x \<noteq> \<omega>"
     let ?g = "\<lambda>x. g' (the_inv_into S f x)"
     show "\<exists>h. simple_function h \<and> (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>) \<and>
               T.simple_integral g' = T.simple_integral (\<lambda>x. h (f x))"
@@ -1068,7 +1068,7 @@
 qed
 
 lemma (in measure_space) positive_integral_vimage_inv:
-  fixes g :: "'d \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  fixes g :: "'d \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a"
   assumes f: "bij_betw f S (space M)"
   shows "measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) g =
       positive_integral (\<lambda>x. g (the_inv_into S f x))"
@@ -1087,8 +1087,8 @@
   and "simple_function u"
   and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
   shows "simple_integral u \<le> (SUP i. positive_integral (f i))" (is "_ \<le> ?S")
-proof (rule pinfreal_le_mult_one_interval)
-  fix a :: pinfreal assume "0 < a" "a < 1"
+proof (rule pextreal_le_mult_one_interval)
+  fix a :: pextreal 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"
@@ -1117,7 +1117,7 @@
     next
       assume "u x \<noteq> 0"
       with `a < 1` real `x \<in> space M`
-      have "a * u x < 1 * u x" by (rule_tac pinfreal_mult_strict_right_mono) (auto simp: image_iff)
+      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
@@ -1130,7 +1130,7 @@
 
   have "simple_integral u = (SUP i. simple_integral (?uB i))"
     unfolding simple_integral_indicator[OF B `simple_function u`]
-  proof (subst SUPR_pinfreal_setsum, safe)
+  proof (subst SUPR_pextreal_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})"
@@ -1142,11 +1142,11 @@
     show "simple_integral 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: pinfreal_SUP_cmult)
+      by (auto intro!: setsum_cong simp: pextreal_SUP_cmult)
   qed
   moreover
   have "a * (SUP i. simple_integral (?uB i)) \<le> ?S"
-    unfolding pinfreal_SUP_cmult[symmetric]
+    unfolding pextreal_SUP_cmult[symmetric]
   proof (safe intro!: SUP_mono bexI)
     fix i
     have "a * simple_integral (?uB i) = simple_integral (\<lambda>x. a * ?uB i x)"
@@ -1306,7 +1306,7 @@
     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_pinfreal_setsum)
+      using insert by (auto intro!: borel_measurable_pextreal_setsum)
     from positive_integral_add[OF this]
     show ?case using insert by auto
   qed simp
@@ -1319,7 +1319,7 @@
   shows "positive_integral (\<lambda>x. f x - g x) = positive_integral f - positive_integral g"
 proof -
   have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-    using f g by (rule borel_measurable_pinfreal_diff)
+    using f g by (rule borel_measurable_pextreal_diff)
   have "positive_integral (\<lambda>x. f x - g x) + positive_integral g =
     positive_integral f"
     unfolding positive_integral_add[OF borel g, symmetric]
@@ -1329,7 +1329,7 @@
       by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
   qed
   with mono show ?thesis
-    by (subst minus_pinfreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
+    by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
 qed
 
 lemma (in measure_space) positive_integral_psuminf:
@@ -1338,7 +1338,7 @@
 proof -
   have "(\<lambda>i. positive_integral (\<lambda>x. \<Sum>i<i. f i x)) \<up> positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity>i. f i x)"
     by (rule positive_integral_isoton)
-       (auto intro!: borel_measurable_pinfreal_setsum assms positive_integral_mono
+       (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
@@ -1347,7 +1347,7 @@
 
 text {* Fatou's lemma: convergence theorem on limes inferior *}
 lemma (in measure_space) positive_integral_lim_INF:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pextreal"
   assumes "\<And>i. u i \<in> borel_measurable M"
   shows "positive_integral (SUP n. INF m. u (m + n)) \<le>
     (SUP n. INF m. positive_integral (u (m + n)))"
@@ -1421,7 +1421,7 @@
   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. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
-    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
+    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times)
   with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
     unfolding isoton_def by simp
 qed
@@ -1493,7 +1493,7 @@
     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!: pinfreal_mult_cancel)
+        by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel)
       finally show "1 \<le> of_nat (Suc n) * u x" .
     qed
     also have "\<dots> = \<mu> ?A"
@@ -1774,7 +1774,7 @@
     using mono by (rule AE_mp) (auto intro!: AE_cong)
   ultimately show ?thesis using fg
     by (auto simp: integral_def integrable_def diff_minus
-             intro!: add_mono real_of_pinfreal_mono positive_integral_mono_AE)
+             intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE)
 qed
 
 lemma (in measure_space) integral_mono:
@@ -1861,7 +1861,7 @@
   also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
     using f by (auto intro!: positive_integral_mono)
   also have "\<dots> < \<omega>"
-    using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
+    using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
   finally have pos: "positive_integral (\<lambda>x. Real (g x)) < \<omega>" .
 
   have "positive_integral (\<lambda>x. Real (- g x)) \<le> positive_integral (\<lambda>x. Real (\<bar>g x\<bar>))"
@@ -1869,7 +1869,7 @@
   also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
     using f by (auto intro!: positive_integral_mono)
   also have "\<dots> < \<omega>"
-    using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
+    using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>)
   finally have neg: "positive_integral (\<lambda>x. Real (- g x)) < \<omega>" .
 
   from neg pos borel show ?thesis
@@ -2018,7 +2018,7 @@
     "positive_integral (\<lambda>x. Real \<bar>f x\<bar>) \<noteq> \<omega>" unfolding integrable_def by auto
   from positive_integral_0_iff[OF this(1)] this(2)
   show ?thesis unfolding integral_def *
-    by (simp add: real_of_pinfreal_eq_0)
+    by (simp add: real_of_pextreal_eq_0)
 qed
 
 lemma (in measure_space) positive_integral_omega:
@@ -2125,8 +2125,8 @@
       by (auto intro!: positive_integral_lim_INF)
     also have "\<dots> = positive_integral (\<lambda>x. Real (2 * w x)) -
         (INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>))"
-      unfolding PI_diff pinfreal_INF_minus[OF I2w_fin] pinfreal_SUP_minus ..
-    finally show ?thesis using neq_0 I2w_fin by (rule pinfreal_le_minus_imp_0)
+      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)
   qed
 
   have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto
@@ -2260,7 +2260,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 (\<lambda>x. \<bar>?F r x\<bar>) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
-      by (simp add: abs_mult_pos real_pinfreal_pos) }
+      by (simp add: abs_mult_pos real_pextreal_pos) }
   note int_abs_F = this
 
   have 1: "\<And>i. integrable (\<lambda>x. ?F i x)"
@@ -2329,8 +2329,8 @@
   show "integrable f" using finite_space finite_measure
     by (simp add: setsum_\<omega> integrable_def)
   show ?I using finite_measure
-    apply (simp add: integral_def real_of_pinfreal_setsum[symmetric]
-      real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric])
+    apply (simp add: integral_def real_of_pextreal_setsum[symmetric]
+      real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric])
     by (rule setsum_cong) (simp_all split: split_if)
 qed
 
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -357,7 +357,7 @@
 qed
 
 lemma lebesgue_simple_function_indicator:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
   assumes f:"lebesgue.simple_function f"
   shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
   apply(rule,subst lebesgue.simple_function_indicator_representation[OF f]) by auto
@@ -421,7 +421,7 @@
 
 lemma lmeasure_singleton[simp]:
   fixes a :: "'a::ordered_euclidean_space" shows "lmeasure {a} = 0"
-  using has_gmeasure_interval[of a a] unfolding zero_pinfreal_def
+  using has_gmeasure_interval[of a a] unfolding zero_pextreal_def
   by (intro has_gmeasure_lmeasure)
      (simp add: content_closed_interval DIM_positive)
 
@@ -475,7 +475,7 @@
 qed
 
 lemma simple_function_has_integral:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
   assumes f:"lebesgue.simple_function f"
   and f':"\<forall>x. f x \<noteq> \<omega>"
   and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
@@ -486,9 +486,9 @@
   have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
     "\<forall>x\<in>range f. x * lmeasure (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
     using f' om unfolding indicator_def by auto
-  show ?case unfolding space_lebesgue real_of_pinfreal_setsum'[OF *(1),THEN sym]
-    unfolding real_of_pinfreal_setsum'[OF *(2),THEN sym]
-    unfolding real_of_pinfreal_setsum space_lebesgue
+  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
@@ -496,8 +496,8 @@
     proof(cases "f y = 0") case False
       have mea:"gmeasurable (f -` {f y})" apply(rule lmeasure_finite_gmeasurable)
         using assms unfolding lebesgue.simple_function_def using False by auto
-      have *:"\<And>x. real (indicator (f -` {f y}) x::pinfreal) = (if x \<in> f -` {f y} then 1 else 0)" by auto
-      show ?thesis unfolding real_of_pinfreal_mult[THEN sym]
+      have *:"\<And>x. real (indicator (f -` {f y}) x::pextreal) = (if x \<in> f -` {f y} then 1 else 0)" by auto
+      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_gmeasure[OF mea,THEN sym]
         unfolding measure_integral_univ[OF mea] * apply(rule integrable_integral)
@@ -510,7 +510,7 @@
   using assms by auto
 
 lemma simple_function_has_integral':
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pextreal"
   assumes f:"lebesgue.simple_function f"
   and i: "lebesgue.simple_integral f \<noteq> \<omega>"
   shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
@@ -544,7 +544,7 @@
 qed
 
 lemma (in measure_space) positive_integral_monotone_convergence:
-  fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  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"
@@ -552,7 +552,7 @@
 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 "(SUP i. f i) = u" using mono lim SUP_Lim_pinfreal
+  have "(SUP i. f i) = u" using mono lim SUP_Lim_pextreal
     unfolding fun_eq_iff SUPR_fun_expand mono_def by auto
   moreover have "(SUP i. f i) \<in> borel_measurable M"
     using i by (rule borel_measurable_SUP)
@@ -560,7 +560,7 @@
 qed
 
 lemma positive_integral_has_integral:
-  fixes f::"'a::ordered_euclidean_space => pinfreal"
+  fixes f::"'a::ordered_euclidean_space => pextreal"
   assumes f:"f \<in> borel_measurable lebesgue"
   and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
   and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
@@ -581,14 +581,14 @@
   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_pinfreal_mono) using u(2,3) by auto
+  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 (lebesgue.positive_integral 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_pinfreal_mono[OF _ int_u_le])
+      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]
 
@@ -921,7 +921,7 @@
 qed
 
 lemma borel_fubini_positiv_integral:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pextreal"
   assumes f: "f \<in> borel_measurable borel"
   shows "borel.positive_integral f =
           borel_product.product_positive_integral {..<DIM('a)} (f \<circ> p2e)"
--- a/src/HOL/Probability/Measure.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Measure.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -103,7 +103,7 @@
     by (rule additiveD [OF additive]) (auto simp add: s)
   finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
   thus ?thesis
-    unfolding minus_pinfreal_eq2[OF s_less_space fin]
+    unfolding minus_pextreal_eq2[OF s_less_space fin]
     by (simp add: ac_simps)
 qed
 
@@ -117,7 +117,7 @@
   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: pinfreal_cancel_plus_minus)
+    by (simp add: pextreal_cancel_plus_minus)
 qed
 
 lemma (in measure_space) measure_countable_increasing:
@@ -225,7 +225,7 @@
     by (rule INF_leI) simp
 
   have "\<mu> (A 0) - (INF n. \<mu> (A n)) = (SUP n. \<mu> (A 0 - A n))"
-    unfolding pinfreal_SUP_minus[symmetric]
+    unfolding pextreal_SUP_minus[symmetric]
     using mono A finite by (subst measure_Diff) auto
   also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
   proof (rule continuity_from_below)
@@ -237,7 +237,7 @@
   also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
     using mono A finite * by (simp, subst measure_Diff) auto
   finally show ?thesis
-    by (rule pinfreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
+    by (rule pextreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
 qed
 
 lemma (in measure_space) measure_down:
@@ -516,7 +516,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 pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
+      by (rule pextreal_less_add[OF not_\<omega>]) (insert contr, auto)
     also have "\<dots> = \<mu> (T \<union> S)"
       using assms by (subst measure_additive) auto
     also have "\<dots> \<le> \<mu> (space M)"
@@ -962,8 +962,8 @@
     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: pinfreal_less_\<omega>)
-    finally show "\<mu> (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pinfreal_less_\<omega>)
+    also have "\<dots> < \<omega>" using `\<mu> (A i) \<noteq> \<omega>` by (auto simp: pextreal_less_\<omega>)
+    finally show "\<mu> (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pextreal_less_\<omega>)
   qed
 qed
 
@@ -1051,14 +1051,14 @@
   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_pinfreal_add)
+  using finite by (auto simp: real_of_pextreal_add)
 
 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>"
   shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-  using real_of_pinfreal_setsum[of S, OF finite]
+  using real_of_pextreal_setsum[of S, OF finite]
         measure_finitely_additive''[symmetric, OF measurable]
   by simp
 
@@ -1093,9 +1093,9 @@
   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_pinfreal_mono)
+    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_pinfreal_add)
+    using fin by (simp add: real_of_pextreal_add)
   finally show ?thesis .
 qed
 
@@ -1104,7 +1104,7 @@
   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_pinfreal_mono measure_countably_subadditive)
+    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 .
@@ -1114,7 +1114,7 @@
   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>"
   shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
-  using measure_finite_singleton[OF S] fin by (simp add: real_of_pinfreal_setsum)
+  using measure_finite_singleton[OF S] fin by (simp add: real_of_pextreal_setsum)
 
 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>"
@@ -1126,7 +1126,7 @@
   note this[simp]
 
   show "mono (\<lambda>i. real (\<mu> (A i)))" unfolding mono_iff_le_Suc using A
-    by (auto intro!: real_of_pinfreal_mono measure_mono)
+    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)]
@@ -1145,7 +1145,7 @@
   note this[simp]
 
   show "mono (\<lambda>i. - real (\<mu> (A i)))" unfolding mono_iff_le_Suc using assms
-    by (auto intro!: real_of_pinfreal_mono measure_mono)
+    by (auto intro!: real_of_pextreal_mono measure_mono)
 
   show "(INF n. Real (real (\<mu> (A n)))) =
     Real (real (\<mu> (\<Inter>i. A i)))"
@@ -1171,8 +1171,8 @@
   hence "\<mu> A \<le> \<mu> (space M)"
     using assms top by (rule measure_mono)
   also have "\<dots> < \<omega>"
-    using finite_measure_of_space unfolding pinfreal_less_\<omega> .
-  finally show ?thesis unfolding pinfreal_less_\<omega> .
+    using finite_measure_of_space unfolding pextreal_less_\<omega> .
+  finally show ?thesis unfolding pextreal_less_\<omega> .
 qed
 
 lemma (in finite_measure) restricted_finite_measure:
@@ -1226,7 +1226,7 @@
 
 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_pinfreal_mono finite_measure)
+  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"
@@ -1449,13 +1449,13 @@
   assumes "disjoint_family A" "x \<in> A i"
   shows "(\<Sum>\<^isub>\<infinity> 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 :: pinfreal)"
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: pextreal)"
     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 :: pinfreal)"
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: pextreal)"
     by (auto simp: setsum_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: pinfreal)"
-  proof (rule pinfreal_SUPI)
-    fix y :: pinfreal 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 :: 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"
     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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Positive_Extended_Real.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -0,0 +1,2775 @@
+(* Author: Johannes Hoelzl, TU Muenchen *)
+
+header {* A type for positive real numbers with infinity *}
+
+theory Positive_Extended_Real
+  imports Complex_Main 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 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 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/Positive_Infinite_Real.thy	Mon Dec 06 19:18:02 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2775 +0,0 @@
-(* Author: Johannes Hoelzl, TU Muenchen *)
-
-header {* A type for positive real numbers with infinity *}
-
-theory Positive_Infinite_Real
-  imports Complex_Main 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 pinfreal = "(Some ` {0::real..}) \<union> {None}"
-  by (rule exI[of _ None]) simp
-
-subsection "Introduce @{typ pinfreal} similar to a datatype"
-
-definition "Real x = Abs_pinfreal (Some (sup 0 x))"
-definition "\<omega> = Abs_pinfreal None"
-
-definition "pinfreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
-
-definition "of_pinfreal = pinfreal_case (\<lambda>x. x) 0"
-
-defs (overloaded)
-  real_of_pinfreal_def [code_unfold]: "real == of_pinfreal"
-
-lemma pinfreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pinfreal"
-  unfolding pinfreal_def by simp
-
-lemma pinfreal_Some_sup[simp]: "Some (sup 0 x) \<in> pinfreal"
-  by (simp add: sup_ge1)
-
-lemma pinfreal_None[simp]: "None \<in> pinfreal"
-  unfolding pinfreal_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_pinfreal_inject)
-
-lemma Real_neq_\<omega>[simp]:
-  "Real x = \<omega> \<longleftrightarrow> False"
-  "\<omega> = Real x \<longleftrightarrow> False"
-  by (simp_all add: Abs_pinfreal_inject \<omega>_def Real_def)
-
-lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
-  unfolding Real_def by (auto simp add: Abs_pinfreal_inject intro!: sup_absorb1)
-
-lemma pinfreal_cases[case_names preal infinite, cases type: pinfreal]:
-  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: pinfreal.Abs_pinfreal_cases)
-  case (Abs_pinfreal y)
-  hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
-    unfolding pinfreal_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_pinfreal(1) sup_absorb2)
-  qed (simp add: \<omega>_def Abs_pinfreal(1) inf)
-qed
-
-lemma pinfreal_case_\<omega>[simp]: "pinfreal_case f i \<omega> = i"
-  unfolding pinfreal_case_def by simp
-
-lemma pinfreal_case_Real[simp]: "pinfreal_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 pinfreal_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 pinfreal_case_def by (simp add: Real_neg)
-qed
-
-lemma pinfreal_case_cancel[simp]: "pinfreal_case (\<lambda>c. i) i x = i"
-  by (cases x) simp_all
-
-lemma pinfreal_case_split:
-  "P (pinfreal_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 pinfreal_case_split_asm:
-  "P (pinfreal_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 pinfreal_case_cong[cong]:
-  assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
-  shows "pinfreal_case f i x = pinfreal_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_pinfreal_def of_pinfreal_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_pinfreal_nonneg[simp, intro]: "0 \<le> real (x :: pinfreal)"
-  unfolding real_of_pinfreal_def of_pinfreal_def
-  by (cases x) auto
-
-lemma real_\<omega>[simp]: "real \<omega> = 0"
-  unfolding real_of_pinfreal_def of_pinfreal_def by simp
-
-lemma pinfreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
-
-subsection "@{typ pinfreal} is a monoid for addition"
-
-instantiation pinfreal :: comm_monoid_add
-begin
-
-definition "0 = Real 0"
-definition "x + y = pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
-
-lemma pinfreal_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_pinfreal_def Real_neg zero_pinfreal_def split: pinfreal_case_split)
-
-lemma \<omega>_neq_0[simp]:
-  "\<omega> = 0 \<longleftrightarrow> False"
-  "0 = \<omega> \<longleftrightarrow> False"
-  by (simp_all add: zero_pinfreal_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_pinfreal_inject zero_pinfreal_def Real_def sup_real_def)
-
-lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pinfreal_def)
-
-instance
-proof
-  fix a :: pinfreal
-  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 pinfreal_plus_eq_\<omega>[simp]: "(a :: pinfreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
-  by (cases a, cases b) auto
-
-lemma pinfreal_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 pinfreal_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_pinfreal_0[simp]: "real (0 :: pinfreal) = 0"
-  unfolding zero_pinfreal_def real_Real by simp
-
-lemma real_of_pinfreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
-  by (cases X) auto
-
-lemma real_of_pinfreal_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_pinfreal_eq_0)
-
-lemma real_of_pinfreal_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: pinfreal_noteq_omega_Ex)
-
-subsection "@{typ pinfreal} is a monoid for multiplication"
-
-instantiation pinfreal :: comm_monoid_mult
-begin
-
-definition "1 = Real 1"
-definition "x * y = (if x = 0 \<or> y = 0 then 0 else
-  pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
-
-lemma pinfreal_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_pinfreal_def one_pinfreal_def)
-
-lemma pinfreal_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_pinfreal_def by simp_all
-
-instance
-proof
-  fix a :: pinfreal show "1 * a = a"
-    by (cases a) (simp_all add: one_pinfreal_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 pinfreal_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 pinfreal_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_pinfreal_def)
-
-lemma real_pinfreal_1[simp]: "real (1 :: pinfreal) = 1"
-  unfolding one_pinfreal_def real_Real by simp
-
-lemma real_of_pinfreal_mult: "real X * real Y = real (X * Y :: pinfreal)"
-  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 pinfreal} is a linear order"
-
-instantiation pinfreal :: linorder
-begin
-
-definition "x < y \<longleftrightarrow> pinfreal_case (\<lambda>i. pinfreal_case (\<lambda>j. i < j) True y) False x"
-definition "x \<le> y \<longleftrightarrow> pinfreal_case (\<lambda>j. pinfreal_case (\<lambda>i. i \<le> j) False x) True y"
-
-lemma pinfreal_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::pinfreal)"
-  by (simp_all add: less_pinfreal_def zero_pinfreal_def one_pinfreal_def del: Real_0 Real_1)
-
-lemma pinfreal_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_pinfreal_def zero_pinfreal_def del: Real_0)
-
-lemma pinfreal_\<omega>_less_eq[simp]:
-  "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
-  by (cases x) (simp_all add: not_le less_eq_pinfreal_def)
-
-lemma pinfreal_less_eq_zero[simp]:
-  "(x::pinfreal) \<le> 0 \<longleftrightarrow> x = 0"
-  by (cases x) (simp_all add: zero_pinfreal_def del: Real_0)
-
-instance
-proof
-  fix x :: pinfreal
-  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 pinfreal_zero_lessI[intro]:
-  "(a :: pinfreal) \<noteq> 0 \<Longrightarrow> 0 < a"
-  by (cases a) auto
-
-lemma pinfreal_less_omegaI[intro, simp]:
-  "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
-  by (cases a) auto
-
-lemma pinfreal_plus_eq_0[simp]: "(a :: pinfreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
-  by (cases a, cases b) auto
-
-lemma pinfreal_le_add1[simp, intro]: "n \<le> n + (m::pinfreal)"
-  by (cases n, cases m) simp_all
-
-lemma pinfreal_le_add2: "(n::pinfreal) + m \<le> k \<Longrightarrow> m \<le> k"
-  by (cases n, cases m, cases k) simp_all
-
-lemma pinfreal_le_add3: "(n::pinfreal) + m \<le> k \<Longrightarrow> n \<le> k"
-  by (cases n, cases m, cases k) simp_all
-
-lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
-  by (cases x) auto
-
-lemma pinfreal_0_less_mult_iff[simp]:
-  fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
-  by (cases x, cases y) (auto simp: zero_less_mult_iff)
-
-lemma pinfreal_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_pinfreal_def del: Real_1)
-
-subsection {* @{text "x - y"} on @{typ pinfreal} *}
-
-instantiation pinfreal :: minus
-begin
-definition "x - y = (if y < x then THE d. x = y + d else 0 :: pinfreal)"
-
-lemma minus_pinfreal_eq:
-  "(x - y = (z :: pinfreal)) \<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_pinfreal_def
-    proof (rule theI2[where Q="\<lambda>d. x = y + d"])
-      show "x = y + pinfreal_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_pinfreal_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_pinfreal_def if_P[OF True] unfolding `x = y + z`
-    proof (rule the_equality)
-      fix d :: pinfreal 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_pinfreal_def)
-qed
-
-instance ..
-end
-
-lemma pinfreal_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::pinfreal) - A = 0"
-  "\<omega> - Real r = \<omega>"
-  "Real r - \<omega> = 0"
-  "A - 0 = A"
-  "0 - A = 0"
-  by (auto simp: minus_pinfreal_eq not_less)
-
-lemma pinfreal_le_epsilon:
-  fixes x y :: pinfreal
-  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 pinfreal :: "{ordered_comm_semiring, comm_semiring_1}"
-proof
-  show "0 \<noteq> (1::pinfreal)" unfolding zero_pinfreal_def one_pinfreal_def
-    by (simp del: Real_1 Real_0)
-
-  fix a :: pinfreal
-  show "0 * a = 0" "a * 0 = 0" by simp_all
-
-  fix b c :: pinfreal
-  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 pinfreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pinfreal) = 0"
-  by (cases x, cases y) (auto simp: mult_le_0_iff)
-
-lemma pinfreal_mult_cancel:
-  fixes x y z :: pinfreal
-  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 pinfreal_of_nat[simp]: "of_nat m = Real (real m)"
-  by (induct m) (auto simp: real_of_nat_Suc one_pinfreal_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_pinfreal_mono:
-  fixes a b :: pinfreal
-  assumes "b \<noteq> \<omega>" "a \<le> b"
-  shows "real a \<le> real b"
-using assms by (cases b, cases a) auto
-
-lemma setprod_pinfreal_0:
-  "(\<Prod>i\<in>I. f i) = (0::pinfreal) \<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_pinfreal_0)
-  qed simp
-qed simp
-
-instance pinfreal :: "semiring_char_0"
-proof
-  fix m n
-  show "inj (of_nat::nat\<Rightarrow>pinfreal)" by (auto intro!: inj_onI)
-qed
-
-subsection "@{typ pinfreal} is a complete lattice"
-
-instantiation pinfreal :: lattice
-begin
-definition [simp]: "sup x y = (max x y :: pinfreal)"
-definition [simp]: "inf x y = (min x y :: pinfreal)"
-instance proof qed simp_all
-end
-
-instantiation pinfreal :: complete_lattice
-begin
-
-definition "bot = Real 0"
-definition "top = \<omega>"
-
-definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pinfreal)"
-definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pinfreal)"
-
-lemma pinfreal_complete_Sup:
-  fixes S :: "pinfreal 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 pinfreal_\<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 pinfreal_complete_Inf:
-  fixes S :: "pinfreal 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 :: pinfreal and A
-
-  show "bot \<le> x" by (cases x) (simp_all add: bot_pinfreal_def)
-  show "x \<le> top" by (simp add: top_pinfreal_def)
-
-  { assume "x \<in> A"
-    with pinfreal_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_pinfreal_def
-      by (auto intro!: Least_equality[symmetric])
-    finally show "x \<le> Sup A" . }
-
-  { assume "x \<in> A"
-    with pinfreal_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_pinfreal_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_pinfreal_def
-        by (auto intro!: Least_equality)
-      thus "Sup A \<le> x" by simp
-    next
-      case False
-      with pinfreal_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_pinfreal_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_pinfreal_def
-        by (auto intro!: Greatest_equality)
-      thus "x \<le> Inf A" by simp
-    next
-      case False
-      with pinfreal_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_pinfreal_def by (auto intro!: Greatest_equality[symmetric])
-      finally show "x \<le> Inf A" .
-    qed }
-qed
-end
-
-lemma Inf_pinfreal_iff:
-  fixes z :: pinfreal
-  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 :: pinfreal assumes "Inf X < z"
-  shows "\<exists>x \<in> X. x < z"
-proof -
-  have "X \<noteq> {}" using assms by (auto simp: Inf_empty top_pinfreal_def)
-  with assms show ?thesis
-    by (metis Inf_pinfreal_iff mem_def not_leE)
-qed
-
-lemma Inf_close:
-  fixes e :: pinfreal 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 pinfreal_SUPI:
-  fixes x :: pinfreal
-  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_pinfreal_def
-  using assms by (auto intro!: Least_equality)
-
-lemma Sup_pinfreal_iff:
-  fixes z :: pinfreal
-  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 :: pinfreal assumes "z < Sup X"
-  shows "\<exists>x \<in> X. z < x"
-proof -
-  have "X \<noteq> {}" using assms by (auto simp: Sup_empty bot_pinfreal_def)
-  with assms show ?thesis
-    by (metis Sup_pinfreal_iff mem_def not_leE)
-qed
-
-lemma Sup_eq_\<omega>: "\<omega> \<in> S \<Longrightarrow> Sup S = \<omega>"
-  unfolding Sup_pinfreal_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: pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_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 pinfreal} *}
-
-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 pinfreal_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 pinfreal_less_eq_diff_eq_sum:
-  fixes x y z :: pinfreal
-  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_pinfreal_eq)
-
-lemma Real_diff_less_omega: "Real r - x < \<omega>" by (cases x) auto
-
-subsubsection {* Numbers on @{typ pinfreal} *}
-
-instantiation pinfreal :: number
-begin
-definition [simp]: "number_of x = Real (number_of x)"
-instance proof qed
-end
-
-subsubsection {* Division on @{typ pinfreal} *}
-
-instantiation pinfreal :: inverse
-begin
-
-definition "inverse x = pinfreal_case (\<lambda>x. if x = 0 then \<omega> else Real (inverse x)) 0 x"
-definition [simp]: "x / y = x * inverse (y :: pinfreal)"
-
-instance proof qed
-end
-
-lemma pinfreal_inverse[simp]:
-  "inverse 0 = \<omega>"
-  "inverse (Real x) = (if x \<le> 0 then \<omega> else Real (inverse x))"
-  "inverse \<omega> = 0"
-  "inverse (1::pinfreal) = 1"
-  "inverse (inverse x) = x"
-  by (simp_all add: inverse_pinfreal_def one_pinfreal_def split: pinfreal_case_split del: Real_1)
-
-lemma pinfreal_inverse_le_eq:
-  assumes "x \<noteq> 0" "x \<noteq> \<omega>"
-  shows "y \<le> z / x \<longleftrightarrow> x * y \<le> (z :: pinfreal)"
-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 :: pinfreal
-  assumes "x < y" shows "inverse y < inverse x"
-  using assms by (cases x, cases y) auto
-
-lemma inverse_antimono:
-  fixes x y :: pinfreal
-  assumes "x \<le> y" shows "inverse y \<le> inverse x"
-  using assms by (cases x, cases y) auto
-
-lemma pinfreal_inverse_\<omega>_iff[simp]: "inverse x = \<omega> \<longleftrightarrow> x = 0"
-  by (cases x) auto
-
-subsection "Infinite sum over @{typ pinfreal}"
-
-text {*
-
-The infinite sum over @{typ pinfreal} has the nice property that it is always
-defined.
-
-*}
-
-definition psuminf :: "(nat \<Rightarrow> pinfreal) \<Rightarrow> pinfreal" (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_pinfreal_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_pinfreal_add setsum_\<omega>)
-qed simp
-
-lemma setsum_0:
-  fixes f :: "'a \<Rightarrow> pinfreal" 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!: pinfreal_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: pinfreal_le_add2)
-  assume "x \<noteq> \<omega>"
-  note move_y = pinfreal_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 = pinfreal_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 pinfreal_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 pinfreal_\<omega>_eq_plus[simp]: "\<omega> = a + b \<longleftrightarrow> (a = \<omega> \<or> b = \<omega>)"
-  by (cases a, cases b) auto
-
-lemma pinfreal_of_nat_le_iff:
-  "(of_nat k :: pinfreal) \<le> of_nat m \<longleftrightarrow> k \<le> m" by auto
-
-lemma pinfreal_of_nat_less_iff:
-  "(of_nat k :: pinfreal) < of_nat m \<longleftrightarrow> k < m" by auto
-
-lemma pinfreal_bound_add:
-  assumes "\<forall>N. f N + y \<le> (x::pinfreal)"
-  shows "(SUP n. f n) + y \<le> x"
-proof (cases "x = \<omega>")
-  have "y \<le> x" using assms by (auto intro: pinfreal_le_add2)
-  assume "x \<noteq> \<omega>"
-  note move_y = pinfreal_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_pinfreal_add:
-  fixes f g :: "nat \<Rightarrow> pinfreal"
-  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 pinfreal_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 pinfreal_bound_add[OF this]
-  have "\<forall>m. (g m) + (SUP n. f n) \<le> y" by (simp add: ac_simps)
-  from pinfreal_bound_add[OF this]
-  show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
-qed
-
-lemma SUPR_pinfreal_setsum:
-  fixes f :: "'x \<Rightarrow> nat \<Rightarrow> pinfreal"
-  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_pinfreal_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_pinfreal_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_pinfreal_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_pinfreal_setsum)
-    by auto
-  finally show ?thesis
-    unfolding psuminf_def by auto
-qed
-
-lemma psuminf_2dimen:
-  fixes f:: "nat * nat \<Rightarrow> pinfreal"
-  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 pinfreal_SUP_cmult:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)"
-proof (rule pinfreal_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 pinfreal_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_pinfreal_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 = pinfreal_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 pinfreal :: 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_pinfreal_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 pinfreal_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_pinfreal_def by auto
-
-instance
-proof
-  let ?U = "UNIV::pinfreal set"
-  show "open ?U" unfolding open_pinfreal_def
-    by (auto intro!: exI[of _ "UNIV"] exI[of _ 0])
-next
-  fix S T::"pinfreal set" assume "open S" and "open T"
-  from `open S`[THEN pinfreal_openE] guess S' xS . note S' = this
-  from `open T`[THEN pinfreal_openE] guess T' xT . note T' = this
-
-  from S'(1-3) T'(1-3)
-  show "open (S \<inter> T)" unfolding open_pinfreal_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:: pinfreal set)"
-  hence "\<forall>S\<in>K. \<exists>T. open T \<and> Real ` (T \<inter> {0..}) = S - {\<omega>}" by (auto simp: open_pinfreal_def)
-  from bchoice[OF this] guess T .. note T = this[rule_format]
-
-  show "open (\<Union>K)" unfolding open_pinfreal_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 pinfreal_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 open_pinfreal_lessThan[simp]:
-  "open {..< a :: pinfreal}"
-proof (cases a)
-  case (preal x) thus ?thesis unfolding open_pinfreal_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_pinfreal_def by (auto intro!: exI[of _ UNIV])
-qed
-
-lemma open_pinfreal_greaterThan[simp]:
-  "open {a :: pinfreal <..}"
-proof (cases a)
-  case (preal x) thus ?thesis unfolding open_pinfreal_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_pinfreal_def by (auto intro!: exI[of _ "{}"])
-qed
-
-lemma pinfreal_open_greaterThanLessThan[simp]: "open {a::pinfreal <..< b}"
-  unfolding greaterThanLessThan_def by auto
-
-lemma closed_pinfreal_atLeast[simp, intro]: "closed {a :: pinfreal ..}"
-proof -
-  have "- {a ..} = {..< a}" by auto
-  then show "closed {a ..}"
-    unfolding closed_def using open_pinfreal_lessThan by auto
-qed
-
-lemma closed_pinfreal_atMost[simp, intro]: "closed {.. b :: pinfreal}"
-proof -
-  have "- {.. b} = {b <..}" by auto
-  then show "closed {.. b}" 
-    unfolding closed_def using open_pinfreal_greaterThan by auto
-qed
-
-lemma closed_pinfreal_atLeastAtMost[simp, intro]:
-  shows "closed {a :: pinfreal .. b}"
-  unfolding atLeastAtMost_def by auto
-
-lemma pinfreal_dense:
-  fixes x y :: pinfreal 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 pinfreal :: t2_space
-proof
-  fix x y :: pinfreal assume "x \<noteq> y"
-  let "?P x (y::pinfreal)" = "\<exists> U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
-
-  { fix x y :: pinfreal assume "x < y"
-    from pinfreal_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::pinfreal)"
-  shows "(\<lambda>i. c * f i) \<up> (c * x)"
-  using assms unfolding isoton_def pinfreal_SUP_cmult
-  by (auto intro: pinfreal_mult_cancel)
-
-lemma isoton_cmult_left:
-  "f \<up> (x::pinfreal) \<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::pinfreal)" and "g \<up> y"
-  shows "(\<lambda>i. f i + g i) \<up> (x + y)"
-  using assms unfolding isoton_def
-  by (auto intro: pinfreal_mult_cancel add_mono simp: SUPR_pinfreal_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 :: pinfreal)"
-  using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
-
-lemma isoton_setsum:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> pinfreal"
-  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 pinfreal_le_mult_one_interval:
-  fixes x y :: pinfreal
-  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 pinfreal_greater_0[intro]:
-  fixes a :: pinfreal
-  assumes "a \<noteq> 0"
-  shows "a > 0"
-using assms apply (cases a) by auto
-
-lemma pinfreal_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 pinfreal_less_\<omega>)
-
-lemma minus_pinfreal_eq2:
-  fixes x y z :: pinfreal
-  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_pinfreal_eq)
-  by (cases x, cases z, auto simp add: ac_simps not_less)
-
-lemma pinfreal_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 pinfreal_inverse_eq_0: "inverse x = 0 \<longleftrightarrow> x = \<omega>"
-  by (cases x) auto
-
-lemma pinfreal_mult_inverse:
-  "\<lbrakk> x \<noteq> \<omega> ; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x * inverse x = 1"
-  by (cases x) auto
-
-lemma pinfreal_zero_less_diff_iff:
-  fixes a b :: pinfreal shows "0 < a - b \<longleftrightarrow> b < a"
-  apply (cases a, cases b)
-  apply (auto simp: pinfreal_noteq_omega_Ex pinfreal_less_\<omega>)
-  apply (cases b)
-  by auto
-
-lemma pinfreal_less_Real_Ex:
-  fixes a b :: pinfreal 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_pinfreal_def apply(rule,rule,rule,rule assms) by auto
-
-lemma pinfreal_zero_le_diff:
-  fixes a b :: pinfreal 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 pinfreal_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 pinfreal_INFI:
-  fixes x :: pinfreal
-  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_pinfreal_def
-  using assms by (auto intro!: Greatest_equality)
-
-lemma real_of_pinfreal_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 pinfreal_less by auto
-qed
-
-lemma not_less_omega[simp]:"\<not> x < \<omega> \<longleftrightarrow> x = \<omega>"
-  by (metis antisym_conv3 pinfreal_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_pinfreal_def of_pinfreal_def
-    unfolding pinfreal_case_def using assms by auto
-qed 
-
-lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)" 
-  unfolding pinfreal_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_pinfreal: assumes inc: "\<And>n m. n\<ge>m \<Longrightarrow> X n \<ge> X m"
-  and lim: "X ----> (L::pinfreal)" 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_pinfreal: assumes "\<And>n m. n\<ge>m \<Longrightarrow> f n \<ge> f m" "f ----> l"
-  shows "(SUP n. f n) = (l::pinfreal)" unfolding SUPR_def Sup_pinfreal_def
-proof (safe intro!: Least_equality)
-  fix n::nat show "f n \<le> l" apply(rule incseq_le_pinfreal)
-    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 pinfreal_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_pinfreal_increasing: assumes "\<forall>n m. n\<ge>m \<longrightarrow> f n \<ge> f m"
-  obtains l where "f ----> (l::pinfreal)"
-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 pinfreal_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_pinfreal_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_pinfreal_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 pinfreal_minus_le_cancel:
-  fixes a b c :: pinfreal
-  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 pinfreal_minus_\<omega>[simp]: "x - \<omega> = 0" by (cases x) simp_all
-
-lemma pinfreal_minus_mono[intro]: "a - x \<le> (a::pinfreal)"
-proof- have "a - x \<le> a - 0"
-    apply(rule pinfreal_minus_le_cancel) by auto
-  thus ?thesis by auto
-qed
-
-lemma pinfreal_minus_eq_\<omega>[simp]: "x - y = \<omega> \<longleftrightarrow> (x = \<omega> \<and> y \<noteq> \<omega>)"
-  by (cases x, cases y) (auto, cases y, auto)
-
-lemma pinfreal_less_minus_iff:
-  fixes a b c :: pinfreal
-  shows "a < b - c \<longleftrightarrow> c + a < b"
-  by (cases c, cases a, cases b, auto)
-
-lemma pinfreal_minus_less_iff:
-  fixes a b c :: pinfreal shows "a - c < b \<longleftrightarrow> (0 < b \<and> (c \<noteq> \<omega> \<longrightarrow> a < b + c))"
-  by (cases c, cases a, cases b, auto)
-
-lemma pinfreal_le_minus_iff:
-  fixes a b c :: pinfreal
-  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: pinfreal_noteq_omega_Ex)
-
-lemma pinfreal_minus_le_iff:
-  fixes a b c :: pinfreal
-  shows "a - c \<le> b \<longleftrightarrow> (c \<le> a \<longrightarrow> a \<le> b + c)"
-  by (cases a, cases c, cases b, auto simp: pinfreal_noteq_omega_Ex)
-
-lemmas pinfreal_minus_order = pinfreal_minus_le_iff pinfreal_minus_less_iff pinfreal_le_minus_iff pinfreal_less_minus_iff
-
-lemma pinfreal_minus_strict_mono:
-  assumes "a > 0" "x > 0" "a\<noteq>\<omega>"
-  shows "a - x < (a::pinfreal)"
-  using assms by(cases x, cases a, auto)
-
-lemma pinfreal_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_pinfreal_eq not_less)
-
-lemma pinfreal_minus_plus:
-  "x \<le> (a::pinfreal) \<Longrightarrow> a - x + x = a"
-  by (cases a, cases x) auto
-
-lemma pinfreal_cancel_plus_minus: "b \<noteq> \<omega> \<Longrightarrow> a + b - b = a"
-  by (cases a, cases b) auto
-
-lemma pinfreal_minus_le_cancel_right:
-  fixes a b c :: pinfreal
-  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_pinfreal_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_pinfreal_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 pinfreal_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!: pinfreal_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: pinfreal_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: pinfreal_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: pinfreal_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::pinfreal)"
-  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: pinfreal_zero_le_diff)
-    have "(a - Sup s) / 2 \<le> a / 2" unfolding divide_pinfreal_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 pinfreal_less using as False
-      by(auto simp add: real_of_pinfreal_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 pinfreal_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_pinfreal_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 pinfreal_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 pinfreal_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: pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_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::pinfreal)"
-  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 pinfreal_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::pinfreal)"
-  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::pinfreal)"
-  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 :: pinfreal
-  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 :: pinfreal
-  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_pinfreal:
-  fixes x :: pinfreal
-  assumes "x > 0" "a + x \<le> b" "a \<noteq> \<omega>"
-  shows "a < b"
-using assms by (cases x, cases a, cases b) auto
-
-lemma pinfreal_INF_minus:
-  fixes f :: "nat \<Rightarrow> pinfreal"
-  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_pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_le_minus_iff by auto
-    show "y \<le> Real x - (SUP i. f i)" unfolding p pinfreal_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 pinfreal_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 pinfreal_le_minus_iff by (auto simp: field_simps)
-      qed
-    qed
-  qed
-qed
-
-lemma pinfreal_SUP_minus:
-  fixes f :: "nat \<Rightarrow> pinfreal"
-  shows "(SUP i. c - f i) = c - (INF i. f i)"
-proof (rule pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_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 pinfreal_le_minus_imp_0:
-  fixes a b :: pinfreal
-  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> pinfreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
-proof -
-  have "\<And>n. 0 < 1 / (of_nat n :: pinfreal)" 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 pinfreal_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 pinfreal_le_epsilon)
-      fix e :: pinfreal 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> pinfreal. 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 pinfreal_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: pinfreal_less_\<omega>)
-      finally have "g i \<noteq> \<omega>" by (simp add: pinfreal_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: pinfreal_noteq_omega_Ex)
-  qed simp
-qed simp
-
-lemma real_of_pinfreal_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_pinfreal_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 pinfreal_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: pinfreal_less_\<omega> pinfreal_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_pinfreal_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_pinfreal_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_pinfreal_inverse_of_nat_Suc_less:
-  fixes e :: pinfreal 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> pinfreal" assumes "mono u"
-  shows "u ----> (SUP i. u i)"
-proof -
-  from lim_pinfreal_increasing[of u] `mono u`
-  obtain l where l: "u ----> l" unfolding mono_def by auto
-  from SUP_Lim_pinfreal[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 :: pinfreal 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> pinfreal"
-  shows "u \<up> x \<longleftrightarrow> (mono u \<and> u ----> x)"
-proof safe
-  assume "mono u" and x: "u ----> x"
-  with SUP_Lim_pinfreal[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 pinfreal_inverse_inverse[simp]:
-  fixes x :: pinfreal
-  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 :: pinfreal .. a} = {.. a}" by auto
-
-lemma greaterThan_omega_Empty: "{\<omega> <..} = {}" by auto
-
-lemma lessThan_0_Empty: "{..< 0 :: pinfreal} = {}" by auto
-
-lemma real_of_pinfreal_inverse[simp]:
-  fixes X :: pinfreal
-  shows "real (inverse X) = 1 / real X"
-  by (cases X) (auto simp: inverse_eq_divide)
-
-lemma real_of_pinfreal_le_0[simp]: "real (X :: pinfreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
-  by (cases X) auto
-
-lemma real_of_pinfreal_less_0[simp]: "\<not> (real (X :: pinfreal) < 0)"
-  by (cases X) auto
-
-lemma abs_real_of_pinfreal[simp]: "\<bar>real (X :: pinfreal)\<bar> = real X"
-  by simp
-
-lemma zero_less_real_of_pinfreal: "0 < real (X :: pinfreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
-  by (cases X) auto
-
-end
--- a/src/HOL/Probability/Probability_Space.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Probability_Space.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -2,24 +2,24 @@
 imports Lebesgue_Integration Radon_Nikodym Product_Measure
 begin
 
-lemma real_of_pinfreal_inverse[simp]:
-  fixes X :: pinfreal
+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_pinfreal_le_0[simp]: "real (X :: pinfreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
+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_pinfreal_less_0[simp]: "\<not> (real (X :: pinfreal) < 0)"
+lemma real_of_pextreal_less_0[simp]: "\<not> (real (X :: pextreal) < 0)"
   by (cases X) auto
 
 locale prob_space = measure_space +
   assumes measure_space_1: "\<mu> (space M) = 1"
 
-lemma abs_real_of_pinfreal[simp]: "\<bar>real (X :: pinfreal)\<bar> = real X"
+lemma abs_real_of_pextreal[simp]: "\<bar>real (X :: pextreal)\<bar> = real X"
   by simp
 
-lemma zero_less_real_of_pinfreal: "0 < real (X :: pinfreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
+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
@@ -141,7 +141,7 @@
   show "prob (\<Union> i :: nat. c i) \<le> 0"
     using real_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_pinfreal_nonneg)
+  show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pextreal_nonneg)
 qed
 
 lemma (in prob_space) indep_sym:
@@ -606,7 +606,7 @@
   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_algebra in_sigma pair_algebraI MX.sets_eq_Pow real_of_pinfreal_setsum)
+    by (simp add: space_pair_algebra in_sigma pair_algebraI MX.sets_eq_Pow real_of_pextreal_setsum)
 qed
 
 lemma (in prob_space) setsum_real_joint_distribution_singleton:
@@ -721,7 +721,7 @@
 
 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_pinfreal_setsum sets_eq_Pow)
+  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>"
@@ -730,27 +730,27 @@
 
 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_pinfreal_pos distribution_finite)
+  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_pinfreal_mult[symmetric]
+  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_pinfreal_def real_of_pinfreal_mult[symmetric]
+  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_pinfreal_def real_of_pinfreal_mult[symmetric]
+  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 prob_space) distribution_remove_const:
@@ -805,9 +805,9 @@
   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_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
-  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
-  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
+  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]:
@@ -821,8 +821,8 @@
 
 lemma (in finite_prob_space) real_distribution_1:
   "real (distribution X A) \<le> 1"
-  unfolding real_pinfreal_1[symmetric]
-  by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp
+  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"
@@ -865,7 +865,7 @@
   unfolding prob_space_def prob_space_axioms_def
 proof
   show "\<mu> (space (restricted_space A)) / \<mu> A = 1"
-    using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pinfreal_noteq_omega_Ex)
+    using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pextreal_noteq_omega_Ex)
   have *: "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
   interpret A: measure_space "restricted_space A" \<mu>
     using `A \<in> sets M` by (rule restricted_measure_space)
@@ -910,9 +910,9 @@
 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_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
-  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
-  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
+  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 finite_prob_space) finite_product_measure_space:
@@ -952,7 +952,7 @@
 section "Conditional Expectation and Probability"
 
 lemma (in prob_space) conditional_expectation_exists:
-  fixes X :: "'a \<Rightarrow> pinfreal"
+  fixes X :: "'a \<Rightarrow> pextreal"
   assumes borel: "X \<in> borel_measurable M"
   and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
   shows "\<exists>Y\<in>borel_measurable (M\<lparr> sets := N \<rparr>). \<forall>C\<in>N.
@@ -999,7 +999,7 @@
   "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
 
 lemma (in prob_space)
-  fixes X :: "'a \<Rightarrow> pinfreal"
+  fixes X :: "'a \<Rightarrow> pextreal"
   assumes borel: "X \<in> borel_measurable M"
   and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
   shows borel_measurable_conditional_expectation:
@@ -1018,7 +1018,7 @@
 qed
 
 lemma (in sigma_algebra) factorize_measurable_function:
-  fixes Z :: "'a \<Rightarrow> pinfreal" and Y :: "'a \<Rightarrow> 'c"
+  fixes Z :: "'a \<Rightarrow> pextreal" 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))"
@@ -1028,7 +1028,7 @@
   from M'.sigma_algebra_vimage[OF this]
   interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 
-  { fix g :: "'c \<Rightarrow> pinfreal" assume "g \<in> borel_measurable M'"
+  { 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)
@@ -1058,7 +1058,7 @@
       show "M'.simple_function ?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::pinfreal)"
+      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)"
         unfolding indicator_def using B by auto
       then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i]
         by (subst va.simple_function_indicator_representation) auto
--- a/src/HOL/Probability/Product_Measure.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Product_Measure.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -379,7 +379,7 @@
       by (auto intro!: M2.finite_measure_compl measurable_cut_fst
                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_pinfreal_diff)
+      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pextreal_diff)
   next
     fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
     moreover then have "\<And>x. \<mu>2 (\<Union>i. Pair x -` F i) = (\<Sum>\<^isub>\<infinity> i. ?s (F i) x)"
@@ -505,7 +505,7 @@
   unfolding pair_measure_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 :: pinfreal)"
+  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: pextreal)"
     unfolding indicator_def by auto
   show "M2.positive_integral (\<lambda>y. indicator A (x, y)) = \<mu>2 (Pair x -` A)"
     unfolding *
@@ -703,7 +703,7 @@
     and "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))
     = positive_integral f"
     by (auto simp del: vimage_Int cong: measurable_cong
-             intro!: M1.borel_measurable_pinfreal_setsum
+             intro!: M1.borel_measurable_pextreal_setsum
              simp add: M1.positive_integral_setsum simple_integral_def
                        M1.positive_integral_cmult
                        M1.positive_integral_cong[OF eq]
@@ -805,7 +805,7 @@
     show "\<mu>1 {x\<in>space M1. \<mu>2 (Pair x -` N) \<noteq> 0} = 0"
       by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def)
     show "{x \<in> space M1. \<mu>2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro M1.borel_measurable_pinfreal_neq_const measure_cut_measurable_fst N)
+      by (intro M1.borel_measurable_pextreal_neq_const measure_cut_measurable_fst N)
     { fix x assume "x \<in> space M1" "\<mu>2 (Pair x -` N) = 0"
       have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
       proof (rule M2.AE_I)
@@ -1201,7 +1201,7 @@
 qed
 
 locale product_sigma_finite =
-  fixes M :: "'i \<Rightarrow> 'a algebra" and \<mu> :: "'i \<Rightarrow> 'a set \<Rightarrow> pinfreal"
+  fixes M :: "'i \<Rightarrow> 'a algebra" and \<mu> :: "'i \<Rightarrow> 'a set \<Rightarrow> pextreal"
   assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i) (\<mu> i)"
 
 locale finite_product_sigma_finite = product_sigma_finite M \<mu> for M :: "'i \<Rightarrow> 'a algebra" and \<mu> +
@@ -1319,7 +1319,7 @@
 qed
 
 definition (in finite_product_sigma_finite)
-  measure :: "('i \<Rightarrow> 'a) set \<Rightarrow> pinfreal" where
+  measure :: "('i \<Rightarrow> 'a) set \<Rightarrow> pextreal" where
   "measure = (SOME \<nu>.
      (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))) \<and>
      sigma_finite_measure P \<nu>)"
--- a/src/HOL/Probability/Radon_Nikodym.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -29,10 +29,10 @@
     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 pinfreal_inverse_eq_0 by simp
+      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: pinfreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
+                 simp: pextreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
     qed
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
@@ -49,7 +49,7 @@
       fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
         using measure[of N] n[of N]
         by (cases "n N")
-           (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff
+           (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)
@@ -65,14 +65,14 @@
     then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
   next
     show "?h \<in> borel_measurable M" using range
-      by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times)
+      by (auto intro!: borel_measurable_psuminf borel_measurable_pextreal_times)
   qed
 qed
 
 subsection "Absolutely continuous"
 
 definition (in measure_space)
-  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))"
+  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pextreal))"
 
 lemma (in sigma_finite_measure) absolutely_continuous_AE:
   assumes "measure_space M \<nu>" "absolutely_continuous \<nu>" "AE x. P x"
@@ -409,9 +409,9 @@
       moreover {
         have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<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: pinfreal_less_\<omega>)
+        also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>)
         finally have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<noteq> \<omega>"
-          by (simp add: pinfreal_less_\<omega>) }
+          by (simp add: pextreal_less_\<omega>) }
       ultimately
       show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
         apply (subst psuminf_minus) by simp_all
@@ -440,7 +440,7 @@
     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: pinfreal_inverse_eq_0 finite_measure_of_space)
+      by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space)
     have "finite_measure M (\<lambda>A. b * \<mu> A)" (is "finite_measure M ?b")
     proof
       show "?b {} = 0" by simp
@@ -486,7 +486,7 @@
         by (cases "positive_integral (\<lambda>x. f x * indicator A x)", cases "\<nu> A", auto)
       finally have "positive_integral (\<lambda>x. ?f0 x * indicator A x) \<le> \<nu> A" . }
     hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
-      by (auto intro!: borel_measurable_indicator borel_measurable_pinfreal_add borel_measurable_pinfreal_times)
+      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>"
       using `A0 \<in> sets M` b
@@ -494,27 +494,27 @@
         finite_measure_of_space M.finite_measure_of_space
       by auto
     have int_f_finite: "positive_integral f \<noteq> \<omega>"
-      using M'.finite_measure_of_space pos_t unfolding pinfreal_zero_less_diff_iff
+      using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_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 pinfreal_mult_inverse)
+      apply (subst pextreal_mult_inverse)
       using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
-      using pinfreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space 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: pinfreal_zero_less_diff_iff)
+      by (simp add: pextreal_zero_less_diff_iff)
     also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
-      using space_less_A0 pos_M pos_t b real[unfolded pinfreal_noteq_omega_Ex] by auto
-    finally have "b * \<mu> A0 < ?t A0" unfolding pinfreal_zero_less_diff_iff .
+      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 `A0 \<in> sets M` by auto
     hence "0 < b * \<mu> A0" using b by auto
     from int_f_finite this
     have "?y + 0 < positive_integral f + b * \<mu> A0" unfolding int_f_eq_y
-      by (rule pinfreal_less_add)
+      by (rule pextreal_less_add)
     also have "\<dots> = positive_integral ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
       by (simp cong: positive_integral_cong)
     finally have "?y < positive_integral ?f0" by simp
@@ -530,7 +530,7 @@
         using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
       show "\<nu> A \<le> positive_integral (\<lambda>x. f x * indicator A x)"
         using upper_bound[THEN bspec, OF `A \<in> sets M`]
-         by (simp add: pinfreal_zero_le_diff)
+         by (simp add: pextreal_zero_le_diff)
     qed
   qed simp
 qed
@@ -573,8 +573,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> pinfreal_less_\<omega> using Q' by auto
-    finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pinfreal_less_\<omega> 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
   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]
@@ -634,7 +634,7 @@
             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: pinfreal_noteq_omega_Ex)
+            by (cases "\<mu> A") (auto simp: pextreal_noteq_omega_Ex)
           with `\<mu> A \<noteq> 0` show ?thesis by auto
         qed
       qed }
@@ -682,7 +682,7 @@
     \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
   proof
     fix i
-    have indicator_eq: "\<And>f x A. (f x :: pinfreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
+    have indicator_eq: "\<And>f x A. (f x :: pextreal) * 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)) \<mu>"
@@ -718,19 +718,19 @@
   show ?thesis
   proof (safe intro!: bexI[of _ ?f])
     show "?f \<in> borel_measurable M"
-      by (safe intro!: borel_measurable_psuminf borel_measurable_pinfreal_times
-        borel_measurable_pinfreal_add borel_measurable_indicator
+      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)
     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 :: pinfreal) =
+      "\<And>x i. (indicator A x * indicator Q0 x :: pextreal) =
         indicator (Q0 \<inter> A) x" by (auto simp: indicator_def)
     have "positive_integral (\<lambda>x. ?f x * indicator A x) =
       (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)"
       unfolding f[OF `A \<in> sets M`]
-      apply (simp del: pinfreal_times(2) add: field_simps *)
+      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)
@@ -775,7 +775,7 @@
   interpret T: finite_measure M ?T
     unfolding finite_measure_def finite_measure_axioms_def
     by (simp 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::pinfreal)} = N"
+  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
@@ -786,10 +786,10 @@
   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_pinfreal_times)
+      using borel f_borel by (auto intro: borel_measurable_pextreal_times)
     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_pinfreal_times borel_measurable_indicator)
+      using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator)
     from positive_integral_translated_density[OF borel this]
     show "\<nu> A = positive_integral (\<lambda>x. h x * f x * indicator A x)"
       unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
@@ -834,7 +834,7 @@
     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: pinfreal_zero_le_diff)
+      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"
@@ -866,15 +866,15 @@
   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::pinfreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+  have *: "\<And>i x A. \<And>y::pextreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
     unfolding indicator_def by auto
   have 1: "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x"
     using borel Q_fin Q
     by (intro finite_density_unique[THEN iffD1] allI)
-       (auto intro!: borel_measurable_pinfreal_times f Int simp: *)
+       (auto intro!: borel_measurable_pextreal_times f Int simp: *)
   have 2: "AE x. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
-    { fix f :: "'a \<Rightarrow> pinfreal" assume borel: "f \<in> borel_measurable M"
+    { fix f :: "'a \<Rightarrow> pextreal" assume borel: "f \<in> borel_measurable M"
         and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
       let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
       have "(\<Union>i. ?A i) \<in> null_sets"
@@ -893,7 +893,7 @@
       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: pinfreal_less_\<omega>) }
+      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" by (simp add: pextreal_less_\<omega>) }
     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)
@@ -927,7 +927,7 @@
   interpret f': measure_space M "\<lambda>A. positive_integral (\<lambda>x. f' x * indicator A x)"
     using borel(2) by (rule measure_space_density)
   { fix A assume "A \<in> sets M"
-    then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pinfreal)} = A"
+    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 "positive_integral (\<lambda>xa. h xa * indicator A xa) = 0 \<longleftrightarrow> A \<in> null_sets"
       using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) }
@@ -1027,7 +1027,7 @@
         apply (subst positive_integral_0_iff)
         apply fast
         apply (subst (asm) AE_iff_null_set)
-        apply (intro borel_measurable_pinfreal_neq_const)
+        apply (intro borel_measurable_pextreal_neq_const)
         apply fast
         by simp
       then show ?thesis by simp
@@ -1130,7 +1130,7 @@
     using sf.RN_deriv(1)[OF \<nu>' ac]
     unfolding measurable_vimage_iff_inv[OF f] comp_def .
   fix A assume "A \<in> sets M"
-  then have *: "\<And>x. x \<in> space M \<Longrightarrow> indicator (f -` A \<inter> S) (the_inv_into S f x) = (indicator A x :: pinfreal)"
+  then have *: "\<And>x. x \<in> space M \<Longrightarrow> indicator (f -` A \<inter> S) (the_inv_into S f x) = (indicator A x :: pextreal)"
     using f[unfolded bij_betw_def]
     unfolding indicator_def by (auto simp: f_the_inv_into_f the_inv_into_in)
   have "\<nu> (f ` (f -` A \<inter> S)) = sf.positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) x * indicator (f -` A \<inter> S) x)"
@@ -1160,7 +1160,7 @@
 proof -
   interpret \<nu>: sigma_finite_measure M \<nu> by fact
   have ms: "measure_space M \<nu>" by default
-  have minus_cong: "\<And>A B A' B'::pinfreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
+  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 f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
   { fix f :: "'a \<Rightarrow> real" assume "f \<in> borel_measurable M"
     { fix x assume *: "RN_deriv \<nu> x \<noteq> \<omega>"
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -8,7 +8,7 @@
   and not_empty[simp]: "S \<noteq> {}"
 
 definition (in finite_space) "M = \<lparr> space = S, sets = Pow S \<rparr>"
-definition (in finite_space) \<mu>_def[simp]: "\<mu> A = (of_nat (card A) / of_nat (card S) :: pinfreal)"
+definition (in finite_space) \<mu>_def[simp]: "\<mu> A = (of_nat (card A) / of_nat (card S) :: pextreal)"
 
 lemma (in finite_space)
   shows space_M[simp]: "space M = S"
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Dec 06 19:18:02 2010 +0100
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Fri Dec 03 15:25:14 2010 +0100
@@ -274,7 +274,7 @@
   "snd ` (SIGMA x:f`X. f -` {x} \<inter> X) = X"
   by (auto simp: image_iff)
 
-lemma zero_le_eq_True: "0 \<le> (x::pinfreal) \<longleftrightarrow> True" by simp
+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"