merged
authorhoelzl
Wed, 20 Jul 2011 10:48:00 +0200
changeset 43925 f651cb053927
parent 43918 6ca79a354c51 (current diff)
parent 43924 1165fe965da8 (diff)
child 43926 3264fbfd87d6
merged
src/HOL/IsaMakefile
src/HOL/Library/Extended_Reals.thy
src/HOL/Library/Nat_Infinity.thy
--- a/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -8,7 +8,7 @@
 imports Buffer Stream_adm
 begin
 
-declare Fin_0 [simp]
+declare enat_0 [simp]
 
 lemma BufAC_Asm_d2: "a\<leadsto>s:BufAC_Asm ==> ? d. a=Md d"
 by (drule BufAC_Asm_unfold [THEN iffD1], auto)
@@ -116,7 +116,7 @@
 done
 
 (*adm_BufAC_Asm,BufAC_Asm_antiton,adm_non_BufAC_Asm,BufAC_Asm_cong*)
-lemma BufAC_Cmt_2stream_monoP: "f:BufEq ==> ? l. !i x s. s:BufAC_Asm --> x << s --> Fin (l i) < #x --> 
+lemma BufAC_Cmt_2stream_monoP: "f:BufEq ==> ? l. !i x s. s:BufAC_Asm --> x << s --> enat (l i) < #x --> 
                      (x,f\<cdot>x):down_iterate BufAC_Cmt_F i --> 
                      (s,f\<cdot>s):down_iterate BufAC_Cmt_F i"
 apply (rule_tac x="%i. 2*i" in exI)
@@ -139,10 +139,10 @@
        \<lbrakk>f \<in> BufEq;
           \<forall>x s. s \<in> BufAC_Asm \<longrightarrow>
                 x \<sqsubseteq> s \<longrightarrow>
-                Fin (2 * i) < #x \<longrightarrow>
+                enat (2 * i) < #x \<longrightarrow>
                 (x, f\<cdot>x) \<in> down_iterate BufAC_Cmt_F i \<longrightarrow>
                 (s, f\<cdot>s) \<in> down_iterate BufAC_Cmt_F i;
-          Md d\<leadsto>\<bullet>\<leadsto>xa \<in> BufAC_Asm; Fin (2 * i) < #ya; f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>t;
+          Md d\<leadsto>\<bullet>\<leadsto>xa \<in> BufAC_Asm; enat (2 * i) < #ya; f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>t;
           (ya, t) \<in> down_iterate BufAC_Cmt_F i; ya \<sqsubseteq> xa\<rbrakk>
        \<Longrightarrow> (xa, rt\<cdot>(f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>xa))) \<in> down_iterate BufAC_Cmt_F i
 *)
@@ -158,11 +158,11 @@
 apply (erule subst)
 (*
  1. \<And>i d xa ya t ff ffa.
-       \<lbrakk>f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>ffa\<cdot>ya; Fin (2 * i) < #ya;
+       \<lbrakk>f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>ffa\<cdot>ya; enat (2 * i) < #ya;
           (ya, ffa\<cdot>ya) \<in> down_iterate BufAC_Cmt_F i; ya \<sqsubseteq> xa; f \<in> BufEq;
           \<forall>x s. s \<in> BufAC_Asm \<longrightarrow>
                 x \<sqsubseteq> s \<longrightarrow>
-                Fin (2 * i) < #x \<longrightarrow>
+                enat (2 * i) < #x \<longrightarrow>
                 (x, f\<cdot>x) \<in> down_iterate BufAC_Cmt_F i \<longrightarrow>
                 (s, f\<cdot>s) \<in> down_iterate BufAC_Cmt_F i;
           xa \<in> BufAC_Asm; ff \<in> BufEq; ffa \<in> BufEq\<rbrakk>
--- a/src/HOL/HOLCF/FOCUS/Fstream.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Fstream.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -144,13 +144,13 @@
 by (simp add: fscons_def)
 
 lemma slen_fscons_eq:
-        "(Fin (Suc n) < #x) = (? a y. x = a~> y & Fin n < #y)"
+        "(enat (Suc n) < #x) = (? a y. x = a~> y & enat n < #y)"
 apply (simp add: fscons_def2 slen_scons_eq)
 apply (fast dest: not_Undef_is_Def [THEN iffD1] elim: DefE)
 done
 
 lemma slen_fscons_eq_rev:
-        "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a~> y | #y < Fin (Suc n))"
+        "(#x < enat (Suc (Suc n))) = (!a y. x ~= a~> y | #y < enat (Suc n))"
 apply (simp add: fscons_def2 slen_scons_eq_rev)
 apply (tactic {* step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1 *})
 apply (tactic {* step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1 *})
@@ -163,7 +163,7 @@
 done
 
 lemma slen_fscons_less_eq:
-        "(#(a~> y) < Fin (Suc (Suc n))) = (#y < Fin (Suc n))"
+        "(#(a~> y) < enat (Suc (Suc n))) = (#y < enat (Suc n))"
 apply (subst slen_fscons_eq_rev)
 apply (fast dest!: fscons_inject [THEN iffD1])
 done
--- a/src/HOL/HOLCF/FOCUS/Fstreams.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Fstreams.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -28,7 +28,7 @@
 
 definition
   jth           :: "nat => 'a fstream => 'a" where
-  "jth = (%n s. if Fin n < #s then THE a. i_th n s = Def a else undefined)"
+  "jth = (%n s. if enat n < #s then THE a. i_th n s = Def a else undefined)"
 
 definition
   first         :: "'a fstream => 'a" where
@@ -36,7 +36,7 @@
 
 definition
   last          :: "'a fstream => 'a" where
-  "last = (%s. case #s of Fin n => (if n~=0 then jth (THE k. Suc k = n) s else undefined))"
+  "last = (%s. case #s of enat n => (if n~=0 then jth (THE k. Suc k = n) s else undefined))"
 
 
 abbreviation
@@ -54,25 +54,25 @@
 lemma ft_fsingleton[simp]: "ft$(<a>) = Def a"
 by (simp add: fsingleton_def2)
 
-lemma slen_fsingleton[simp]: "#(<a>) = Fin 1"
-by (simp add: fsingleton_def2 inat_defs)
+lemma slen_fsingleton[simp]: "#(<a>) = enat 1"
+by (simp add: fsingleton_def2 enat_defs)
 
 lemma slen_fstreams[simp]: "#(<a> ooo s) = iSuc (#s)"
 by (simp add: fsingleton_def2)
 
 lemma slen_fstreams2[simp]: "#(s ooo <a>) = iSuc (#s)"
 apply (cases "#s")
-apply (auto simp add: iSuc_Fin)
+apply (auto simp add: iSuc_enat)
 apply (insert slen_sconc [of _ s "Suc 0" "<a>"], auto)
 by (simp add: sconc_def)
 
 lemma j_th_0_fsingleton[simp]:"jth 0 (<a>) = a"
 apply (simp add: fsingleton_def2 jth_def)
-by (simp add: i_th_def Fin_0)
+by (simp add: i_th_def enat_0)
 
 lemma jth_0[simp]: "jth 0 (<a> ooo s) = a"  
 apply (simp add: fsingleton_def2 jth_def)
-by (simp add: i_th_def Fin_0)
+by (simp add: i_th_def enat_0)
 
 lemma first_sconc[simp]: "first (<a> ooo s) = a"
 by (simp add: first_def)
@@ -80,14 +80,14 @@
 lemma first_fsingleton[simp]: "first (<a>) = a"
 by (simp add: first_def)
 
-lemma jth_n[simp]: "Fin n = #s ==> jth n (s ooo <a>) = a"
+lemma jth_n[simp]: "enat n = #s ==> jth n (s ooo <a>) = a"
 apply (simp add: jth_def, auto)
 apply (simp add: i_th_def rt_sconc1)
-by (simp add: inat_defs split: inat.splits)
+by (simp add: enat_defs split: enat.splits)
 
-lemma last_sconc[simp]: "Fin n = #s ==> last (s ooo <a>) = a"
+lemma last_sconc[simp]: "enat n = #s ==> last (s ooo <a>) = a"
 apply (simp add: last_def)
-apply (simp add: inat_defs split:inat.splits)
+apply (simp add: enat_defs split:enat.splits)
 by (drule sym, auto)
 
 lemma last_fsingleton[simp]: "last (<a>) = a"
@@ -97,18 +97,18 @@
 by (simp add: first_def jth_def)
 
 lemma last_UU[simp]:"last UU = undefined"
-by (simp add: last_def jth_def inat_defs)
+by (simp add: last_def jth_def enat_defs)
 
-lemma last_infinite[simp]:"#s = Infty ==> last s = undefined"
+lemma last_infinite[simp]:"#s = \<infinity> ==> last s = undefined"
 by (simp add: last_def)
 
-lemma jth_slen_lemma1:"n <= k & Fin n = #s ==> jth k s = undefined"
-by (simp add: jth_def inat_defs split:inat.splits, auto)
+lemma jth_slen_lemma1:"n <= k & enat n = #s ==> jth k s = undefined"
+by (simp add: jth_def enat_defs split:enat.splits, auto)
 
 lemma jth_UU[simp]:"jth n UU = undefined" 
 by (simp add: jth_def)
 
-lemma ext_last:"[|s ~= UU; Fin (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" 
+lemma ext_last:"[|s ~= UU; enat (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" 
 apply (simp add: last_def)
 apply (case_tac "#s", auto)
 apply (simp add: fsingleton_def2)
--- a/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -10,14 +10,14 @@
 
 definition
   stream_monoP  :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
-  "stream_monoP F = (\<exists>Q i. \<forall>P s. Fin i \<le> #s \<longrightarrow>
+  "stream_monoP F = (\<exists>Q i. \<forall>P s. enat i \<le> #s \<longrightarrow>
                     (s \<in> F P) = (stream_take i\<cdot>s \<in> Q \<and> iterate i\<cdot>rt\<cdot>s \<in> P))"
 
 definition
   stream_antiP  :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
   "stream_antiP F = (\<forall>P x. \<exists>Q i.
-                (#x  < Fin i \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow> y \<in> F P \<longrightarrow> x \<in> F P)) \<and>
-                (Fin i <= #x \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow>
+                (#x  < enat i \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow> y \<in> F P \<longrightarrow> x \<in> F P)) \<and>
+                (enat i <= #x \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow>
                 (y \<in> F P) = (stream_take i\<cdot>y \<in> Q \<and> iterate i\<cdot>rt\<cdot>y \<in> P))))"
 
 definition
@@ -57,7 +57,7 @@
 lemma flatstream_adm_lemma:
   assumes 1: "Porder.chain Y"
   assumes 2: "!i. P (Y i)"
-  assumes 3: "(!!Y. [| Porder.chain Y; !i. P (Y i); !k. ? j. Fin k < #((Y j)::'a::flat stream)|]
+  assumes 3: "(!!Y. [| Porder.chain Y; !i. P (Y i); !k. ? j. enat k < #((Y j)::'a::flat stream)|]
   ==> P(LUB i. Y i))"
   shows "P(LUB i. Y i)"
 apply (rule increasing_chain_adm_lemma [of _ P, OF 1 2])
@@ -74,12 +74,12 @@
 apply (   erule spec)
 apply (  assumption)
 apply ( assumption)
-apply (metis inat_ord_code(4) slen_infinite)
+apply (metis enat_ord_code(4) slen_infinite)
 done
 
 (* should be without reference to stream length? *)
 lemma flatstream_admI: "[|(!!Y. [| Porder.chain Y; !i. P (Y i); 
- !k. ? j. Fin k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P"
+ !k. ? j. enat k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P"
 apply (unfold adm_def)
 apply (intro strip)
 apply (erule (1) flatstream_adm_lemma)
@@ -87,13 +87,13 @@
 done
 
 
-(* context (theory "Nat_InFinity");*)
-lemma ile_lemma: "Fin (i + j) <= x ==> Fin i <= x"
+(* context (theory "Extended_Nat");*)
+lemma ile_lemma: "enat (i + j) <= x ==> enat i <= x"
   by (rule order_trans) auto
 
 lemma stream_monoP2I:
 "!!X. stream_monoP F ==> !i. ? l. !x y. 
-  Fin l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i"
+  enat l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i"
 apply (unfold stream_monoP_def)
 apply (safe)
 apply (rule_tac x="i*ia" in exI)
@@ -120,7 +120,7 @@
 done
 
 lemma stream_monoP2_gfp_admI: "[| !i. ? l. !x y. 
- Fin l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i;
+ enat l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i;
     down_cont F |] ==> adm (%x. x:gfp F)"
 apply (erule INTER_down_iterate_is_gfp [THEN ssubst]) (* cont *)
 apply (simp (no_asm))
@@ -153,7 +153,7 @@
 apply (intro strip)
 apply (erule allE, erule all_dupE, erule exE, erule exE)
 apply (erule conjE)
-apply (case_tac "#x < Fin i")
+apply (case_tac "#x < enat i")
 apply ( fast)
 apply (unfold linorder_not_less)
 apply (drule (1) mp)
--- a/src/HOL/HOLCF/IsaMakefile	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/IsaMakefile	Wed Jul 20 10:48:00 2011 +0200
@@ -134,7 +134,7 @@
 HOLCF-ex: HOLCF $(LOG)/HOLCF-ex.gz
 
 $(LOG)/HOLCF-ex.gz: $(OUT)/HOLCF \
-  ../Library/Nat_Infinity.thy \
+  ../Library/Extended_Nat.thy \
   ex/Concurrency_Monad.thy \
   ex/Dagstuhl.thy \
   ex/Dnat.thy \
--- a/src/HOL/HOLCF/Library/Stream.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/HOLCF/Library/Stream.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -5,7 +5,7 @@
 header {* General Stream domain *}
 
 theory Stream
-imports HOLCF "~~/src/HOL/Library/Nat_Infinity"
+imports HOLCF "~~/src/HOL/Library/Extended_Nat"
 begin
 
 default_sort pcpo
@@ -22,8 +22,8 @@
                                      If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)"
 
 definition
-  slen :: "'a stream \<Rightarrow> inat"  ("#_" [1000] 1000) where
-  "#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
+  slen :: "'a stream \<Rightarrow> enat"  ("#_" [1000] 1000) where
+  "#s = (if stream_finite s then enat (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
 
 
 (* concatenation *)
@@ -39,7 +39,7 @@
 definition
   sconc :: "'a stream => 'a stream => 'a stream"  (infixr "ooo" 65) where
   "s1 ooo s2 = (case #s1 of
-                  Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
+                  enat n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
                | \<infinity>     \<Rightarrow> s1)"
 
 primrec constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream"
@@ -51,7 +51,7 @@
 definition
   constr_sconc  :: "'a stream => 'a stream => 'a stream" where (* constructive *)
   "constr_sconc s1 s2 = (case #s1 of
-                          Fin n \<Rightarrow> constr_sconc' n s1 s2
+                          enat n \<Rightarrow> constr_sconc' n s1 s2
                         | \<infinity>    \<Rightarrow> s1)"
 
 
@@ -327,12 +327,12 @@
 section "slen"
 
 lemma slen_empty [simp]: "#\<bottom> = 0"
-by (simp add: slen_def stream.finite_def zero_inat_def Least_equality)
+by (simp add: slen_def stream.finite_def zero_enat_def Least_equality)
 
 lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = iSuc (#xs)"
 apply (case_tac "stream_finite (x && xs)")
 apply (simp add: slen_def, auto)
-apply (simp add: stream.finite_def, auto simp add: iSuc_Fin)
+apply (simp add: stream.finite_def, auto simp add: iSuc_enat)
 apply (rule Least_Suc2, auto)
 (*apply (drule sym)*)
 (*apply (drule sym scons_eq_UU [THEN iffD1],simp)*)
@@ -340,13 +340,13 @@
 apply (simp add: slen_def, auto)
 by (drule stream_finite_lemma1,auto)
 
-lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \<bottom>)"
-by (cases x, auto simp add: Fin_0 iSuc_Fin[THEN sym])
+lemma slen_less_1_eq: "(#x < enat (Suc 0)) = (x = \<bottom>)"
+by (cases x, auto simp add: enat_0 iSuc_enat[THEN sym])
 
 lemma slen_empty_eq: "(#x = 0) = (x = \<bottom>)"
 by (cases x, auto)
 
-lemma slen_scons_eq: "(Fin (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  Fin n < #y)"
+lemma slen_scons_eq: "(enat (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  enat n < #y)"
 apply (auto, case_tac "x=UU",auto)
 apply (drule stream_exhaust_eq [THEN iffD1], auto)
 apply (case_tac "#y") apply simp_all
@@ -359,18 +359,18 @@
 lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \<infinity>"
 by (simp add: slen_def)
 
-lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < Fin (Suc n))"
+lemma slen_scons_eq_rev: "(#x < enat (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < enat (Suc n))"
  apply (cases x, auto)
-   apply (simp add: zero_inat_def)
-  apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
- apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
+   apply (simp add: zero_enat_def)
+  apply (case_tac "#stream") apply (simp_all add: iSuc_enat)
+ apply (case_tac "#stream") apply (simp_all add: iSuc_enat)
 done
 
 lemma slen_take_lemma4 [rule_format]:
-  "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n"
-apply (induct n, auto simp add: Fin_0)
+  "!s. stream_take n$s ~= s --> #(stream_take n$s) = enat n"
+apply (induct n, auto simp add: enat_0)
 apply (case_tac "s=UU", simp)
-by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_Fin)
+by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_enat)
 
 (*
 lemma stream_take_idempotent [simp]:
@@ -390,41 +390,41 @@
 by (simp add: chain_def,simp)
 *)
 
-lemma slen_take_eq: "ALL x. (Fin n < #x) = (stream_take n\<cdot>x ~= x)"
+lemma slen_take_eq: "ALL x. (enat n < #x) = (stream_take n\<cdot>x ~= x)"
 apply (induct_tac n, auto)
-apply (simp add: Fin_0, clarsimp)
+apply (simp add: enat_0, clarsimp)
 apply (drule not_sym)
 apply (drule slen_empty_eq [THEN iffD1], simp)
 apply (case_tac "x=UU", simp)
 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
 apply (erule_tac x="y" in allE, auto)
-apply (simp_all add: not_less iSuc_Fin)
+apply (simp_all add: not_less iSuc_enat)
 apply (case_tac "#y") apply simp_all
 apply (case_tac "x=UU", simp)
 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
 apply (erule_tac x="y" in allE, simp)
 apply (case_tac "#y") by simp_all
 
-lemma slen_take_eq_rev: "(#x <= Fin n) = (stream_take n\<cdot>x = x)"
+lemma slen_take_eq_rev: "(#x <= enat n) = (stream_take n\<cdot>x = x)"
 by (simp add: linorder_not_less [symmetric] slen_take_eq)
 
-lemma slen_take_lemma1: "#x = Fin n ==> stream_take n\<cdot>x = x"
+lemma slen_take_lemma1: "#x = enat n ==> stream_take n\<cdot>x = x"
 by (rule slen_take_eq_rev [THEN iffD1], auto)
 
 lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)"
 apply (cases s1)
  by (cases s2, simp+)+
 
-lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"
+lemma slen_take_lemma5: "#(stream_take n$s) <= enat n"
 apply (case_tac "stream_take n$s = s")
  apply (simp add: slen_take_eq_rev)
 by (simp add: slen_take_lemma4)
 
-lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = Fin i"
+lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = enat i"
 apply (simp add: stream.finite_def, auto)
 by (simp add: slen_take_lemma4)
 
-lemma slen_infinite: "stream_finite x = (#x ~= Infty)"
+lemma slen_infinite: "stream_finite x = (#x ~= \<infinity>)"
 by (simp add: slen_def)
 
 lemma slen_mono_lemma: "stream_finite s ==> ALL t. s << t --> #s <= #t"
@@ -443,19 +443,19 @@
 lemma iterate_lemma: "F$(iterate n$F$x) = iterate n$F$(F$x)"
 by (insert iterate_Suc2 [of n F x], auto)
 
-lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i$rt$x)"
+lemma slen_rt_mult [rule_format]: "!x. enat (i + j) <= #x --> enat j <= #(iterate i$rt$x)"
 apply (induct i, auto)
-apply (case_tac "x=UU", auto simp add: zero_inat_def)
+apply (case_tac "x=UU", auto simp add: zero_enat_def)
 apply (drule stream_exhaust_eq [THEN iffD1], auto)
 apply (erule_tac x="y" in allE, auto)
-apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_Fin)
+apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_enat)
 by (simp add: iterate_lemma)
 
 lemma slen_take_lemma3 [rule_format]:
-  "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
+  "!(x::'a::flat stream) y. enat n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
 apply (induct_tac n, auto)
 apply (case_tac "x=UU", auto)
-apply (simp add: zero_inat_def)
+apply (simp add: zero_enat_def)
 apply (simp add: Suc_ile_eq)
 apply (case_tac "y=UU", clarsimp)
 apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+
@@ -478,7 +478,7 @@
 apply (subgoal_tac "stream_take n$s ~=s")
  apply (insert slen_take_lemma4 [of n s],auto)
 apply (cases s, simp)
-by (simp add: slen_take_lemma4 iSuc_Fin)
+by (simp add: slen_take_lemma4 iSuc_enat)
 
 (* ----------------------------------------------------------------------- *)
 (* theorems about smap                                                     *)
@@ -546,7 +546,7 @@
 lemma i_rt_mono: "x << s ==> i_rt n x  << i_rt n s"
 by (simp add: i_rt_def monofun_rt_mult)
 
-lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)"
+lemma i_rt_ij_lemma: "enat (i + j) <= #x ==> enat j <= #(i_rt i x)"
 by (simp add: i_rt_def slen_rt_mult)
 
 lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)"
@@ -566,14 +566,14 @@
  apply (subgoal_tac "#(i_rt n s)=0")
   apply (case_tac "stream_take n$s = s",simp+)
   apply (insert slen_take_eq [rule_format,of n s],simp)
-  apply (cases "#s") apply (simp_all add: zero_inat_def)
+  apply (cases "#s") apply (simp_all add: zero_enat_def)
   apply (simp add: slen_take_eq)
   apply (cases "#s")
   using i_rt_take_lemma1 [of n s]
-  apply (simp_all add: zero_inat_def)
+  apply (simp_all add: zero_enat_def)
   done
 
-lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
+lemma i_rt_lemma_slen: "#s=enat n ==> i_rt n s = UU"
 by (simp add: i_rt_slen slen_take_lemma1)
 
 lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s"
@@ -581,29 +581,29 @@
  apply (cases s, auto simp del: i_rt_Suc)
 by (simp add: i_rt_Suc_back stream_finite_rt_eq)+
 
-lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl &
-                            #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j
-                                              --> Fin (j + t) = #x"
+lemma take_i_rt_len_lemma: "ALL sl x j t. enat sl = #x & n <= sl &
+                            #(stream_take n$x) = enat t & #(i_rt n x)= enat j
+                                              --> enat (j + t) = #x"
 apply (induct n, auto)
- apply (simp add: zero_inat_def)
+ apply (simp add: zero_enat_def)
 apply (case_tac "x=UU",auto)
- apply (simp add: zero_inat_def)
+ apply (simp add: zero_enat_def)
 apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
-apply (subgoal_tac "EX k. Fin k = #y",clarify)
+apply (subgoal_tac "EX k. enat k = #y",clarify)
  apply (erule_tac x="k" in allE)
  apply (erule_tac x="y" in allE,auto)
  apply (erule_tac x="THE p. Suc p = t" in allE,auto)
-   apply (simp add: iSuc_def split: inat.splits)
-  apply (simp add: iSuc_def split: inat.splits)
+   apply (simp add: iSuc_def split: enat.splits)
+  apply (simp add: iSuc_def split: enat.splits)
   apply (simp only: the_equality)
- apply (simp add: iSuc_def split: inat.splits)
+ apply (simp add: iSuc_def split: enat.splits)
  apply force
-apply (simp add: iSuc_def split: inat.splits)
+apply (simp add: iSuc_def split: enat.splits)
 done
 
 lemma take_i_rt_len:
-"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==>
-    Fin (j + t) = #x"
+"[| enat sl = #x; n <= sl; #(stream_take n$x) = enat t; #(i_rt n x) = enat j |] ==>
+    enat (j + t) = #x"
 by (blast intro: take_i_rt_len_lemma [rule_format])
 
 
@@ -690,13 +690,13 @@
 (* ----------------------------------------------------------------------- *)
 
 lemma UU_sconc [simp]: " UU ooo s = s "
-by (simp add: sconc_def zero_inat_def)
+by (simp add: sconc_def zero_enat_def)
 
 lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
 by auto
 
 lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
-apply (simp add: sconc_def zero_inat_def iSuc_def split: inat.splits, auto)
+apply (simp add: sconc_def zero_enat_def iSuc_def split: enat.splits, auto)
 apply (rule someI2_ex,auto)
  apply (rule_tac x="x && y" in exI,auto)
 apply (simp add: i_rt_Suc_forw)
@@ -704,21 +704,21 @@
 by (drule stream_exhaust_eq [THEN iffD1],auto)
 
 lemma ex_sconc [rule_format]:
-  "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)"
+  "ALL k y. #x = enat k --> (EX w. stream_take k$w = x & i_rt k w = y)"
 apply (case_tac "#x")
  apply (rule stream_finite_ind [of x],auto)
   apply (simp add: stream.finite_def)
   apply (drule slen_take_lemma1,blast)
- apply (simp_all add: zero_inat_def iSuc_def split: inat.splits)
+ apply (simp_all add: zero_enat_def iSuc_def split: enat.splits)
 apply (erule_tac x="y" in allE,auto)
 by (rule_tac x="a && w" in exI,auto)
 
-lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"
-apply (simp add: sconc_def split: inat.splits, arith?,auto)
+lemma rt_sconc1: "enat n = #x ==> i_rt n (x ooo y) = y"
+apply (simp add: sconc_def split: enat.splits, arith?,auto)
 apply (rule someI2_ex,auto)
 by (drule ex_sconc,simp)
 
-lemma sconc_inj2: "\<lbrakk>Fin n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
+lemma sconc_inj2: "\<lbrakk>enat n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
 apply (frule_tac y=y in rt_sconc1)
 by (auto elim: rt_sconc1)
 
@@ -734,7 +734,7 @@
  apply (simp add: i_rt_slen)
 by (simp add: sconc_def)
 
-lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x"
+lemma stream_take_sconc [simp]: "enat n = #x ==> stream_take n$(x ooo y) = x"
 apply (simp add: sconc_def)
 apply (cases "#x")
 apply auto
@@ -743,7 +743,7 @@
 
 lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y"
 apply (cases "#x",auto)
- apply (simp add: sconc_def iSuc_Fin)
+ apply (simp add: sconc_def iSuc_enat)
  apply (rule someI2_ex)
   apply (drule ex_sconc, simp)
  apply (rule someI2_ex, auto)
@@ -799,9 +799,9 @@
 by (cases s, auto)
 
 lemma i_th_sconc_lemma [rule_format]:
-  "ALL x y. Fin n < #x --> i_th n (x ooo y) = i_th n x"
+  "ALL x y. enat n < #x --> i_th n (x ooo y) = i_th n x"
 apply (induct_tac n, auto)
-apply (simp add: Fin_0 i_th_def)
+apply (simp add: enat_0 i_th_def)
 apply (simp add: slen_empty_eq ft_sconc)
 apply (simp add: i_th_def)
 apply (case_tac "x=UU",auto)
@@ -849,16 +849,16 @@
 (* ----------------------------------------------------------------------- *)
 
 lemma slen_sconc_finite1:
-  "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty"
-apply (case_tac "#y ~= Infty",auto)
+  "[| #(x ooo y) = \<infinity>; enat n = #x |] ==> #y = \<infinity>"
+apply (case_tac "#y ~= \<infinity>",auto)
 apply (drule_tac y=y in rt_sconc1)
 apply (insert stream_finite_i_rt [of n "x ooo y"])
 by (simp add: slen_infinite)
 
-lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty"
+lemma slen_sconc_infinite1: "#x=\<infinity> ==> #(x ooo y) = \<infinity>"
 by (simp add: sconc_def)
 
-lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty"
+lemma slen_sconc_infinite2: "#y=\<infinity> ==> #(x ooo y) = \<infinity>"
 apply (case_tac "#x")
  apply (simp add: sconc_def)
  apply (rule someI2_ex)
@@ -868,7 +868,7 @@
  apply (fastsimp simp add: slen_infinite,auto)
 by (simp add: sconc_def)
 
-lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)"
+lemma sconc_finite: "(#x~=\<infinity> & #y~=\<infinity>) = (#(x ooo y)~=\<infinity>)"
 apply auto
   apply (metis not_Infty_eq slen_sconc_finite1)
  apply (metis not_Infty_eq slen_sconc_infinite1)
@@ -877,7 +877,7 @@
 
 (* ----------------------------------------------------------------------- *)
 
-lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k"
+lemma slen_sconc_mono3: "[| enat n = #x; enat k = #(x ooo y) |] ==> n <= k"
 apply (insert slen_mono [of "x" "x ooo y"])
 apply (cases "#x") apply simp_all
 apply (cases "#(x ooo y)") apply simp_all
@@ -887,10 +887,10 @@
    subsection "finite slen"
 (* ----------------------------------------------------------------------- *)
 
-lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)"
+lemma slen_sconc: "[| enat n = #x; enat m = #y |] ==> #(x ooo y) = enat (n + m)"
 apply (case_tac "#(x ooo y)")
  apply (frule_tac y=y in rt_sconc1)
- apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp)
+ apply (insert take_i_rt_len [of "THE j. enat j = #(x ooo y)" "x ooo y" n n m],simp)
  apply (insert slen_sconc_mono3 [of n x _ y],simp)
 by (insert sconc_finite [of x y],auto)
 
@@ -934,7 +934,7 @@
 
 lemma contlub_sconc_lemma:
   "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
-apply (case_tac "#x=Infty")
+apply (case_tac "#x=\<infinity>")
  apply (simp add: sconc_def)
 apply (drule finite_lub_sconc,auto simp add: slen_infinite)
 done
@@ -948,7 +948,7 @@
 (* ----------------------------------------------------------------------- *)
 
 lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
-by (simp add: constr_sconc_def zero_inat_def)
+by (simp add: constr_sconc_def zero_enat_def)
 
 lemma "x ooo y = constr_sconc x y"
 apply (case_tac "#x")
@@ -956,7 +956,7 @@
   defer 1
   apply (simp add: constr_sconc_def del: scons_sconc)
   apply (case_tac "#s")
-   apply (simp add: iSuc_Fin)
+   apply (simp add: iSuc_enat)
    apply (case_tac "a=UU",auto simp del: scons_sconc)
    apply (simp)
   apply (simp add: sconc_def)
--- a/src/HOL/IsaMakefile	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/IsaMakefile	Wed Jul 20 10:48:00 2011 +0200
@@ -444,25 +444,25 @@
   Library/AssocList.thy Library/BigO.thy Library/Binomial.thy		\
   Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy	\
   Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy	\
-  Library/Code_Char_ord.thy Library/Code_Integer.thy Library/Code_Natural.thy			\
-  Library/Code_Prolog.thy Tools/Predicate_Compile/code_prolog.ML	\
-  Library/ContNotDenum.thy Library/Continuity.thy Library/Convex.thy	\
-  Library/Countable.thy Library/Diagonalize.thy Library/Dlist.thy	\
-  Library/Dlist_Cset.thy 						\
+  Library/Code_Char_ord.thy Library/Code_Integer.thy			\
+  Library/Code_Natural.thy Library/Code_Prolog.thy			\
+  Tools/Predicate_Compile/code_prolog.ML Library/ContNotDenum.thy	\
+  Library/Continuity.thy Library/Convex.thy Library/Countable.thy	\
+  Library/Diagonalize.thy Library/Dlist.thy Library/Dlist_Cset.thy 	\
   Library/Efficient_Nat.thy Library/Eval_Witness.thy 			\
-  Library/Executable_Set.thy Library/Extended_Reals.thy			\
-  Library/Float.thy Library/Formal_Power_Series.thy			\
-  Library/Fraction_Field.thy Library/FrechetDeriv.thy Library/Cset.thy	\
-  Library/FuncSet.thy Library/Function_Algebras.thy			\
-  Library/Fundamental_Theorem_Algebra.thy Library/Glbs.thy		\
-  Library/Indicator_Function.thy Library/Infinite_Set.thy		\
-  Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
-  Library/LaTeXsugar.thy Library/Lattice_Algebras.thy			\
-  Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy 	\
-  Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy	\
-  Library/Monad_Syntax.thy Library/More_List.thy Library/More_Set.thy	\
-  Library/Multiset.thy Library/Nat_Bijection.thy			\
-  Library/Nat_Infinity.thy Library/Nested_Environment.thy		\
+  Library/Executable_Set.thy Library/Extended_Real.thy			\
+  Library/Extended_Nat.thy Library/Float.thy				\
+  Library/Formal_Power_Series.thy Library/Fraction_Field.thy		\
+  Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy		\
+  Library/Function_Algebras.thy Library/Fundamental_Theorem_Algebra.thy	\
+  Library/Glbs.thy Library/Indicator_Function.thy			\
+  Library/Infinite_Set.thy Library/Inner_Product.thy			\
+  Library/Kleene_Algebra.thy Library/LaTeXsugar.thy			\
+  Library/Lattice_Algebras.thy Library/Lattice_Syntax.thy		\
+  Library/Library.thy Library/List_Cset.thy Library/List_Prefix.thy	\
+  Library/List_lexord.thy Library/Mapping.thy Library/Monad_Syntax.thy	\
+  Library/More_List.thy Library/More_Set.thy Library/Multiset.thy	\
+  Library/Nat_Bijection.thy Library/Nested_Environment.thy		\
   Library/Numeral_Type.thy Library/OptionalSugar.thy			\
   Library/Order_Relation.thy Library/Permutation.thy			\
   Library/Permutations.thy Library/Poly_Deriv.thy			\
@@ -481,7 +481,7 @@
   Library/Sum_of_Squares/sos_wrapper.ML					\
   Library/Sum_of_Squares/sum_of_squares.ML				\
   Library/Transitive_Closure_Table.thy Library/Univ_Poly.thy		\
-  Library/While_Combinator.thy Library/Zorn.thy	\
+  Library/While_Combinator.thy Library/Zorn.thy				\
   $(SRC)/Tools/adhoc_overloading.ML Library/positivstellensatz.ML	\
   Library/reflection.ML Library/reify_data.ML				\
   Library/document/root.bib Library/document/root.tex
@@ -1201,7 +1201,7 @@
   Multivariate_Analysis/Topology_Euclidean_Space.thy			\
   Multivariate_Analysis/document/root.tex				\
   Multivariate_Analysis/normarith.ML Library/Glbs.thy			\
-  Library/Extended_Reals.thy Library/Indicator_Function.thy		\
+  Library/Extended_Real.thy Library/Indicator_Function.thy		\
   Library/Inner_Product.thy Library/Numeral_Type.thy Library/Convex.thy	\
   Library/FrechetDeriv.thy Library/Product_Vector.thy			\
   Library/Product_plus.thy
@@ -1574,7 +1574,7 @@
 HOLCF-Library: HOLCF $(LOG)/HOLCF-Library.gz
 
 $(LOG)/HOLCF-Library.gz: $(OUT)/HOLCF \
-  Library/Nat_Infinity.thy \
+  Library/Extended_Nat.thy \
   HOLCF/Library/Defl_Bifinite.thy \
   HOLCF/Library/Bool_Discrete.thy \
   HOLCF/Library/Char_Discrete.thy \
@@ -1628,7 +1628,7 @@
 HOLCF-FOCUS: HOLCF $(LOG)/HOLCF-FOCUS.gz
 
 $(LOG)/HOLCF-FOCUS.gz: $(OUT)/HOLCF HOLCF/FOCUS/ROOT.ML \
-  Library/Nat_Infinity.thy \
+  Library/Extended_Nat.thy \
   HOLCF/Library/Stream.thy \
   HOLCF/FOCUS/Fstreams.thy \
   HOLCF/FOCUS/Fstream.thy \
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Extended_Nat.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -0,0 +1,573 @@
+(*  Title:      HOL/Library/Extended_Nat.thy
+    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
+    Contributions: David Trachtenherz, TU Muenchen
+*)
+
+header {* Extended natural numbers (i.e. with infinity) *}
+
+theory Extended_Nat
+imports Main
+begin
+
+class infinity =
+  fixes infinity :: "'a"
+
+notation (xsymbols)
+  infinity  ("\<infinity>")
+
+notation (HTML output)
+  infinity  ("\<infinity>")
+
+subsection {* Type definition *}
+
+text {*
+  We extend the standard natural numbers by a special value indicating
+  infinity.
+*}
+
+typedef (open) enat = "UNIV :: nat option set" ..
+ 
+definition enat :: "nat \<Rightarrow> enat" where
+  "enat n = Abs_enat (Some n)"
+ 
+instantiation enat :: infinity
+begin
+  definition "\<infinity> = Abs_enat None"
+  instance proof qed
+end
+ 
+rep_datatype enat "\<infinity> :: enat"
+proof -
+  fix P i assume "\<And>j. P (enat j)" "P \<infinity>"
+  then show "P i"
+  proof induct
+    case (Abs_enat y) then show ?case
+      by (cases y rule: option.exhaust)
+         (auto simp: enat_def infinity_enat_def)
+  qed
+qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject)
+
+declare [[coercion "enat::nat\<Rightarrow>enat"]]
+
+lemma not_Infty_eq[iff]: "(x \<noteq> \<infinity>) = (EX i. x = enat i)"
+by (cases x) auto
+
+lemma not_enat_eq [iff]: "(ALL y. x ~= enat y) = (x = \<infinity>)"
+by (cases x) auto
+
+primrec the_enat :: "enat \<Rightarrow> nat"
+where "the_enat (enat n) = n"
+
+subsection {* Constructors and numbers *}
+
+instantiation enat :: "{zero, one, number}"
+begin
+
+definition
+  "0 = enat 0"
+
+definition
+  [code_unfold]: "1 = enat 1"
+
+definition
+  [code_unfold, code del]: "number_of k = enat (number_of k)"
+
+instance ..
+
+end
+
+definition iSuc :: "enat \<Rightarrow> enat" where
+  "iSuc i = (case i of enat n \<Rightarrow> enat (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
+
+lemma enat_0: "enat 0 = 0"
+  by (simp add: zero_enat_def)
+
+lemma enat_1: "enat 1 = 1"
+  by (simp add: one_enat_def)
+
+lemma enat_number: "enat (number_of k) = number_of k"
+  by (simp add: number_of_enat_def)
+
+lemma one_iSuc: "1 = iSuc 0"
+  by (simp add: zero_enat_def one_enat_def iSuc_def)
+
+lemma Infty_ne_i0 [simp]: "(\<infinity>::enat) \<noteq> 0"
+  by (simp add: zero_enat_def)
+
+lemma i0_ne_Infty [simp]: "0 \<noteq> (\<infinity>::enat)"
+  by (simp add: zero_enat_def)
+
+lemma zero_enat_eq [simp]:
+  "number_of k = (0\<Colon>enat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
+  "(0\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
+  unfolding zero_enat_def number_of_enat_def by simp_all
+
+lemma one_enat_eq [simp]:
+  "number_of k = (1\<Colon>enat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
+  "(1\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
+  unfolding one_enat_def number_of_enat_def by simp_all
+
+lemma zero_one_enat_neq [simp]:
+  "\<not> 0 = (1\<Colon>enat)"
+  "\<not> 1 = (0\<Colon>enat)"
+  unfolding zero_enat_def one_enat_def by simp_all
+
+lemma Infty_ne_i1 [simp]: "(\<infinity>::enat) \<noteq> 1"
+  by (simp add: one_enat_def)
+
+lemma i1_ne_Infty [simp]: "1 \<noteq> (\<infinity>::enat)"
+  by (simp add: one_enat_def)
+
+lemma Infty_ne_number [simp]: "(\<infinity>::enat) \<noteq> number_of k"
+  by (simp add: number_of_enat_def)
+
+lemma number_ne_Infty [simp]: "number_of k \<noteq> (\<infinity>::enat)"
+  by (simp add: number_of_enat_def)
+
+lemma iSuc_enat: "iSuc (enat n) = enat (Suc n)"
+  by (simp add: iSuc_def)
+
+lemma iSuc_number_of: "iSuc (number_of k) = enat (Suc (number_of k))"
+  by (simp add: iSuc_enat number_of_enat_def)
+
+lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
+  by (simp add: iSuc_def)
+
+lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
+  by (simp add: iSuc_def zero_enat_def split: enat.splits)
+
+lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
+  by (rule iSuc_ne_0 [symmetric])
+
+lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
+  by (simp add: iSuc_def split: enat.splits)
+
+lemma number_of_enat_inject [simp]:
+  "(number_of k \<Colon> enat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
+  by (simp add: number_of_enat_def)
+
+
+subsection {* Addition *}
+
+instantiation enat :: comm_monoid_add
+begin
+
+definition [nitpick_simp]:
+  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | enat m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | enat n \<Rightarrow> enat (m + n)))"
+
+lemma plus_enat_simps [simp, code]:
+  fixes q :: enat
+  shows "enat m + enat n = enat (m + n)"
+    and "\<infinity> + q = \<infinity>"
+    and "q + \<infinity> = \<infinity>"
+  by (simp_all add: plus_enat_def split: enat.splits)
+
+instance proof
+  fix n m q :: enat
+  show "n + m + q = n + (m + q)"
+    by (cases n, auto, cases m, auto, cases q, auto)
+  show "n + m = m + n"
+    by (cases n, auto, cases m, auto)
+  show "0 + n = n"
+    by (cases n) (simp_all add: zero_enat_def)
+qed
+
+end
+
+lemma plus_enat_0 [simp]:
+  "0 + (q\<Colon>enat) = q"
+  "(q\<Colon>enat) + 0 = q"
+  by (simp_all add: plus_enat_def zero_enat_def split: enat.splits)
+
+lemma plus_enat_number [simp]:
+  "(number_of k \<Colon> enat) + number_of l = (if k < Int.Pls then number_of l
+    else if l < Int.Pls then number_of k else number_of (k + l))"
+  unfolding number_of_enat_def plus_enat_simps nat_arith(1) if_distrib [symmetric, of _ enat] ..
+
+lemma iSuc_number [simp]:
+  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
+  unfolding iSuc_number_of
+  unfolding one_enat_def number_of_enat_def Suc_nat_number_of if_distrib [symmetric] ..
+
+lemma iSuc_plus_1:
+  "iSuc n = n + 1"
+  by (cases n) (simp_all add: iSuc_enat one_enat_def)
+  
+lemma plus_1_iSuc:
+  "1 + q = iSuc q"
+  "q + 1 = iSuc q"
+by (simp_all add: iSuc_plus_1 add_ac)
+
+lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
+by (simp_all add: iSuc_plus_1 add_ac)
+
+lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
+by (simp only: add_commute[of m] iadd_Suc)
+
+lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)"
+by (cases m, cases n, simp_all add: zero_enat_def)
+
+subsection {* Multiplication *}
+
+instantiation enat :: comm_semiring_1
+begin
+
+definition times_enat_def [nitpick_simp]:
+  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | enat m \<Rightarrow>
+    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | enat n \<Rightarrow> enat (m * n)))"
+
+lemma times_enat_simps [simp, code]:
+  "enat m * enat n = enat (m * n)"
+  "\<infinity> * \<infinity> = (\<infinity>::enat)"
+  "\<infinity> * enat n = (if n = 0 then 0 else \<infinity>)"
+  "enat m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
+  unfolding times_enat_def zero_enat_def
+  by (simp_all split: enat.split)
+
+instance proof
+  fix a b c :: enat
+  show "(a * b) * c = a * (b * c)"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "a * b = b * a"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "1 * a = a"
+    unfolding times_enat_def zero_enat_def one_enat_def
+    by (simp split: enat.split)
+  show "(a + b) * c = a * c + b * c"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split add: left_distrib)
+  show "0 * a = 0"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "a * 0 = 0"
+    unfolding times_enat_def zero_enat_def
+    by (simp split: enat.split)
+  show "(0::enat) \<noteq> 1"
+    unfolding zero_enat_def one_enat_def
+    by simp
+qed
+
+end
+
+lemma mult_iSuc: "iSuc m * n = n + m * n"
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
+
+lemma mult_iSuc_right: "m * iSuc n = m + m * n"
+  unfolding iSuc_plus_1 by (simp add: algebra_simps)
+
+lemma of_nat_eq_enat: "of_nat n = enat n"
+  apply (induct n)
+  apply (simp add: enat_0)
+  apply (simp add: plus_1_iSuc iSuc_enat)
+  done
+
+instance enat :: number_semiring
+proof
+  fix n show "number_of (int n) = (of_nat n :: enat)"
+    unfolding number_of_enat_def number_of_int of_nat_id of_nat_eq_enat ..
+qed
+
+instance enat :: semiring_char_0 proof
+  have "inj enat" by (rule injI) simp
+  then show "inj (\<lambda>n. of_nat n :: enat)" by (simp add: of_nat_eq_enat)
+qed
+
+lemma imult_is_0[simp]: "((m::enat) * n = 0) = (m = 0 \<or> n = 0)"
+by(auto simp add: times_enat_def zero_enat_def split: enat.split)
+
+lemma imult_is_Infty: "((a::enat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
+by(auto simp add: times_enat_def zero_enat_def split: enat.split)
+
+
+subsection {* Subtraction *}
+
+instantiation enat :: minus
+begin
+
+definition diff_enat_def:
+"a - b = (case a of (enat x) \<Rightarrow> (case b of (enat y) \<Rightarrow> enat (x - y) | \<infinity> \<Rightarrow> 0)
+          | \<infinity> \<Rightarrow> \<infinity>)"
+
+instance ..
+
+end
+
+lemma idiff_enat_enat[simp,code]: "enat a - enat b = enat (a - b)"
+by(simp add: diff_enat_def)
+
+lemma idiff_Infty[simp,code]: "\<infinity> - n = (\<infinity>::enat)"
+by(simp add: diff_enat_def)
+
+lemma idiff_Infty_right[simp,code]: "enat a - \<infinity> = 0"
+by(simp add: diff_enat_def)
+
+lemma idiff_0[simp]: "(0::enat) - n = 0"
+by (cases n, simp_all add: zero_enat_def)
+
+lemmas idiff_enat_0[simp] = idiff_0[unfolded zero_enat_def]
+
+lemma idiff_0_right[simp]: "(n::enat) - 0 = n"
+by (cases n) (simp_all add: zero_enat_def)
+
+lemmas idiff_enat_0_right[simp] = idiff_0_right[unfolded zero_enat_def]
+
+lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::enat) - n = 0"
+by(auto simp: zero_enat_def)
+
+lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
+by(simp add: iSuc_def split: enat.split)
+
+lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
+by(simp add: one_enat_def iSuc_enat[symmetric] zero_enat_def[symmetric])
+
+(*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*)
+
+subsection {* Ordering *}
+
+instantiation enat :: linordered_ab_semigroup_add
+begin
+
+definition [nitpick_simp]:
+  "m \<le> n = (case n of enat n1 \<Rightarrow> (case m of enat m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
+    | \<infinity> \<Rightarrow> True)"
+
+definition [nitpick_simp]:
+  "m < n = (case m of enat m1 \<Rightarrow> (case n of enat n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
+    | \<infinity> \<Rightarrow> False)"
+
+lemma enat_ord_simps [simp]:
+  "enat m \<le> enat n \<longleftrightarrow> m \<le> n"
+  "enat m < enat n \<longleftrightarrow> m < n"
+  "q \<le> (\<infinity>::enat)"
+  "q < (\<infinity>::enat) \<longleftrightarrow> q \<noteq> \<infinity>"
+  "(\<infinity>::enat) \<le> q \<longleftrightarrow> q = \<infinity>"
+  "(\<infinity>::enat) < q \<longleftrightarrow> False"
+  by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits)
+
+lemma enat_ord_code [code]:
+  "enat m \<le> enat n \<longleftrightarrow> m \<le> n"
+  "enat m < enat n \<longleftrightarrow> m < n"
+  "q \<le> (\<infinity>::enat) \<longleftrightarrow> True"
+  "enat m < \<infinity> \<longleftrightarrow> True"
+  "\<infinity> \<le> enat n \<longleftrightarrow> False"
+  "(\<infinity>::enat) < q \<longleftrightarrow> False"
+  by simp_all
+
+instance by default
+  (auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits)
+
+end
+
+instance enat :: ordered_comm_semiring
+proof
+  fix a b c :: enat
+  assume "a \<le> b" and "0 \<le> c"
+  thus "c * a \<le> c * b"
+    unfolding times_enat_def less_eq_enat_def zero_enat_def
+    by (simp split: enat.splits)
+qed
+
+lemma enat_ord_number [simp]:
+  "(number_of m \<Colon> enat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
+  "(number_of m \<Colon> enat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
+  by (simp_all add: number_of_enat_def)
+
+lemma i0_lb [simp]: "(0\<Colon>enat) \<le> n"
+  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma ile0_eq [simp]: "n \<le> (0\<Colon>enat) \<longleftrightarrow> n = 0"
+by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma Infty_ileE [elim!]: "\<infinity> \<le> enat m \<Longrightarrow> R"
+  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
+
+lemma Infty_ilessE [elim!]: "\<infinity> < enat m \<Longrightarrow> R"
+  by simp
+
+lemma not_iless0 [simp]: "\<not> n < (0\<Colon>enat)"
+  by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma i0_less [simp]: "(0\<Colon>enat) < n \<longleftrightarrow> n \<noteq> 0"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
+  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
+ 
+lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
+  by (simp add: iSuc_def less_enat_def split: enat.splits)
+
+lemma ile_iSuc [simp]: "n \<le> iSuc n"
+  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
+
+lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
+  by (simp add: zero_enat_def iSuc_def less_eq_enat_def split: enat.splits)
+
+lemma i0_iless_iSuc [simp]: "0 < iSuc n"
+  by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.splits)
+
+lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
+by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.split)
+
+lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
+  by (simp add: iSuc_def less_eq_enat_def less_enat_def split: enat.splits)
+
+lemma Suc_ile_eq: "enat (Suc m) \<le> n \<longleftrightarrow> enat m < n"
+  by (cases n) auto
+
+lemma iless_Suc_eq [simp]: "enat m < iSuc n \<longleftrightarrow> enat m \<le> n"
+  by (auto simp add: iSuc_def less_enat_def split: enat.splits)
+
+lemma imult_Infty: "(0::enat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma imult_Infty_right: "(0::enat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
+by (simp add: zero_enat_def less_enat_def split: enat.splits)
+
+lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \<and> 0 < n)"
+by (simp only: i0_less imult_is_0, simp)
+
+lemma mono_iSuc: "mono iSuc"
+by(simp add: mono_def)
+
+
+lemma min_enat_simps [simp]:
+  "min (enat m) (enat n) = enat (min m n)"
+  "min q 0 = 0"
+  "min 0 q = 0"
+  "min q (\<infinity>::enat) = q"
+  "min (\<infinity>::enat) q = q"
+  by (auto simp add: min_def)
+
+lemma max_enat_simps [simp]:
+  "max (enat m) (enat n) = enat (max m n)"
+  "max q 0 = q"
+  "max 0 q = q"
+  "max q \<infinity> = (\<infinity>::enat)"
+  "max \<infinity> q = (\<infinity>::enat)"
+  by (simp_all add: max_def)
+
+lemma enat_ile: "n \<le> enat m \<Longrightarrow> \<exists>k. n = enat k"
+  by (cases n) simp_all
+
+lemma enat_iless: "n < enat m \<Longrightarrow> \<exists>k. n = enat k"
+  by (cases n) simp_all
+
+lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. enat k < Y j"
+apply (induct_tac k)
+ apply (simp (no_asm) only: enat_0)
+ apply (fast intro: le_less_trans [OF i0_lb])
+apply (erule exE)
+apply (drule spec)
+apply (erule exE)
+apply (drule ileI1)
+apply (rule iSuc_enat [THEN subst])
+apply (rule exI)
+apply (erule (1) le_less_trans)
+done
+
+instantiation enat :: "{bot, top}"
+begin
+
+definition bot_enat :: enat where
+  "bot_enat = 0"
+
+definition top_enat :: enat where
+  "top_enat = \<infinity>"
+
+instance proof
+qed (simp_all add: bot_enat_def top_enat_def)
+
+end
+
+lemma finite_enat_bounded:
+  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> enat n"
+  shows "finite A"
+proof (rule finite_subset)
+  show "finite (enat ` {..n})" by blast
+
+  have "A \<subseteq> {..enat n}" using le_fin by fastsimp
+  also have "\<dots> \<subseteq> enat ` {..n}"
+    by (rule subsetI) (case_tac x, auto)
+  finally show "A \<subseteq> enat ` {..n}" .
+qed
+
+
+subsection {* Well-ordering *}
+
+lemma less_enatE:
+  "[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P"
+by (induct n) auto
+
+lemma less_InftyE:
+  "[| n < \<infinity>; !!k. n = enat k ==> P |] ==> P"
+by (induct n) auto
+
+lemma enat_less_induct:
+  assumes prem: "!!n. \<forall>m::enat. m < n --> P m ==> P n" shows "P n"
+proof -
+  have P_enat: "!!k. P (enat k)"
+    apply (rule nat_less_induct)
+    apply (rule prem, clarify)
+    apply (erule less_enatE, simp)
+    done
+  show ?thesis
+  proof (induct n)
+    fix nat
+    show "P (enat nat)" by (rule P_enat)
+  next
+    show "P \<infinity>"
+      apply (rule prem, clarify)
+      apply (erule less_InftyE)
+      apply (simp add: P_enat)
+      done
+  qed
+qed
+
+instance enat :: wellorder
+proof
+  fix P and n
+  assume hyp: "(\<And>n\<Colon>enat. (\<And>m\<Colon>enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
+  show "P n" by (blast intro: enat_less_induct hyp)
+qed
+
+subsection {* Complete Lattice *}
+
+instantiation enat :: complete_lattice
+begin
+
+definition inf_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
+  "inf_enat \<equiv> min"
+
+definition sup_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
+  "sup_enat \<equiv> max"
+
+definition Inf_enat :: "enat set \<Rightarrow> enat" where
+  "Inf_enat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
+
+definition Sup_enat :: "enat set \<Rightarrow> enat" where
+  "Sup_enat A \<equiv> if A = {} then 0
+    else if finite A then Max A
+                     else \<infinity>"
+instance proof
+  fix x :: "enat" and A :: "enat set"
+  { assume "x \<in> A" then show "Inf A \<le> x"
+      unfolding Inf_enat_def by (auto intro: Least_le) }
+  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
+      unfolding Inf_enat_def
+      by (cases "A = {}") (auto intro: LeastI2_ex) }
+  { assume "x \<in> A" then show "x \<le> Sup A"
+      unfolding Sup_enat_def by (cases "finite A") auto }
+  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
+      unfolding Sup_enat_def using finite_enat_bounded by auto }
+qed (simp_all add: inf_enat_def sup_enat_def)
+end
+
+
+subsection {* Traditional theorem names *}
+
+lemmas enat_defs = zero_enat_def one_enat_def number_of_enat_def iSuc_def
+  plus_enat_def less_eq_enat_def less_enat_def
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Extended_Real.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -0,0 +1,2543 @@
+(*  Title:      HOL/Library/Extended_Real.thy
+    Author:     Johannes Hölzl, TU München
+    Author:     Robert Himmelmann, TU München
+    Author:     Armin Heller, TU München
+    Author:     Bogdan Grechuk, University of Edinburgh
+*)
+
+header {* Extended real number line *}
+
+theory Extended_Real
+  imports Complex_Main Extended_Nat
+begin
+
+text {*
+
+For more lemmas about the extended real numbers go to
+  @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"}
+
+*}
+
+lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
+proof
+  assume "{x..} = UNIV"
+  show "x = bot"
+  proof (rule ccontr)
+    assume "x \<noteq> bot" then have "bot \<notin> {x..}" by (simp add: le_less)
+    then show False using `{x..} = UNIV` by simp
+  qed
+qed auto
+
+lemma SUPR_pair:
+  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
+  by (rule antisym) (auto intro!: SUP_leI le_SUPI2)
+
+lemma INFI_pair:
+  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
+  by (rule antisym) (auto intro!: le_INFI INF_leI2)
+
+subsection {* Definition and basic properties *}
+
+datatype ereal = ereal real | PInfty | MInfty
+
+instantiation ereal :: uminus
+begin
+  fun uminus_ereal where
+    "- (ereal r) = ereal (- r)"
+  | "- PInfty = MInfty"
+  | "- MInfty = PInfty"
+  instance ..
+end
+
+instantiation ereal :: infinity
+begin
+  definition "(\<infinity>::ereal) = PInfty"
+  instance ..
+end
+
+definition "ereal_of_enat n = (case n of enat n \<Rightarrow> ereal (real n) | \<infinity> \<Rightarrow> \<infinity>)"
+
+declare [[coercion "ereal :: real \<Rightarrow> ereal"]]
+declare [[coercion "ereal_of_enat :: enat \<Rightarrow> ereal"]]
+declare [[coercion "(\<lambda>n. ereal (of_nat n)) :: nat \<Rightarrow> ereal"]]
+
+lemma ereal_uminus_uminus[simp]:
+  fixes a :: ereal shows "- (- a) = a"
+  by (cases a) simp_all
+
+lemma
+  shows PInfty_eq_infinity[simp]: "PInfty = \<infinity>"
+    and MInfty_eq_minfinity[simp]: "MInfty = - \<infinity>"
+    and MInfty_neq_PInfty[simp]: "\<infinity> \<noteq> - (\<infinity>::ereal)" "- \<infinity> \<noteq> (\<infinity>::ereal)"
+    and MInfty_neq_ereal[simp]: "ereal r \<noteq> - \<infinity>" "- \<infinity> \<noteq> ereal r"
+    and PInfty_neq_ereal[simp]: "ereal r \<noteq> \<infinity>" "\<infinity> \<noteq> ereal r"
+    and PInfty_cases[simp]: "(case \<infinity> of ereal r \<Rightarrow> f r | PInfty \<Rightarrow> y | MInfty \<Rightarrow> z) = y"
+    and MInfty_cases[simp]: "(case - \<infinity> of ereal r \<Rightarrow> f r | PInfty \<Rightarrow> y | MInfty \<Rightarrow> z) = z"
+  by (simp_all add: infinity_ereal_def)
+
+lemma inj_ereal[simp]: "inj_on ereal A"
+  unfolding inj_on_def by auto
+
+lemma ereal_cases[case_names real PInf MInf, cases type: ereal]:
+  assumes "\<And>r. x = ereal r \<Longrightarrow> P"
+  assumes "x = \<infinity> \<Longrightarrow> P"
+  assumes "x = -\<infinity> \<Longrightarrow> P"
+  shows P
+  using assms by (cases x) auto
+
+lemmas ereal2_cases = ereal_cases[case_product ereal_cases]
+lemmas ereal3_cases = ereal2_cases[case_product ereal_cases]
+
+lemma ereal_uminus_eq_iff[simp]:
+  fixes a b :: ereal shows "-a = -b \<longleftrightarrow> a = b"
+  by (cases rule: ereal2_cases[of a b]) simp_all
+
+function of_ereal :: "ereal \<Rightarrow> real" where
+"of_ereal (ereal r) = r" |
+"of_ereal \<infinity> = 0" |
+"of_ereal (-\<infinity>) = 0"
+  by (auto intro: ereal_cases)
+termination proof qed (rule wf_empty)
+
+defs (overloaded)
+  real_of_ereal_def [code_unfold]: "real \<equiv> of_ereal"
+
+lemma real_of_ereal[simp]:
+    "real (- x :: ereal) = - (real x)"
+    "real (ereal r) = r"
+    "real (\<infinity>::ereal) = 0"
+  by (cases x) (simp_all add: real_of_ereal_def)
+
+lemma range_ereal[simp]: "range ereal = UNIV - {\<infinity>, -\<infinity>}"
+proof safe
+  fix x assume "x \<notin> range ereal" "x \<noteq> \<infinity>"
+  then show "x = -\<infinity>" by (cases x) auto
+qed auto
+
+lemma ereal_range_uminus[simp]: "range uminus = (UNIV::ereal set)"
+proof safe
+  fix x :: ereal show "x \<in> range uminus" by (intro image_eqI[of _ _ "-x"]) auto
+qed auto
+
+instantiation ereal :: number
+begin
+definition [simp]: "number_of x = ereal (number_of x)"
+instance proof qed
+end
+
+instantiation ereal :: abs
+begin
+  function abs_ereal where
+    "\<bar>ereal r\<bar> = ereal \<bar>r\<bar>"
+  | "\<bar>-\<infinity>\<bar> = (\<infinity>::ereal)"
+  | "\<bar>\<infinity>\<bar> = (\<infinity>::ereal)"
+  by (auto intro: ereal_cases)
+  termination proof qed (rule wf_empty)
+  instance ..
+end
+
+lemma abs_eq_infinity_cases[elim!]: "\<lbrakk> \<bar>x :: ereal\<bar> = \<infinity> ; x = \<infinity> \<Longrightarrow> P ; x = -\<infinity> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+  by (cases x) auto
+
+lemma abs_neq_infinity_cases[elim!]: "\<lbrakk> \<bar>x :: ereal\<bar> \<noteq> \<infinity> ; \<And>r. x = ereal r \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+  by (cases x) auto
+
+lemma abs_ereal_uminus[simp]: "\<bar>- x\<bar> = \<bar>x::ereal\<bar>"
+  by (cases x) auto
+
+subsubsection "Addition"
+
+instantiation ereal :: comm_monoid_add
+begin
+
+definition "0 = ereal 0"
+
+function plus_ereal where
+"ereal r + ereal p = ereal (r + p)" |
+"\<infinity> + a = (\<infinity>::ereal)" |
+"a + \<infinity> = (\<infinity>::ereal)" |
+"ereal r + -\<infinity> = - \<infinity>" |
+"-\<infinity> + ereal p = -(\<infinity>::ereal)" |
+"-\<infinity> + -\<infinity> = -(\<infinity>::ereal)"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a, b)" by (cases x) auto
+  ultimately show P
+   by (cases rule: ereal2_cases[of a b]) auto
+qed auto
+termination proof qed (rule wf_empty)
+
+lemma Infty_neq_0[simp]:
+  "(\<infinity>::ereal) \<noteq> 0" "0 \<noteq> (\<infinity>::ereal)"
+  "-(\<infinity>::ereal) \<noteq> 0" "0 \<noteq> -(\<infinity>::ereal)"
+  by (simp_all add: zero_ereal_def)
+
+lemma ereal_eq_0[simp]:
+  "ereal r = 0 \<longleftrightarrow> r = 0"
+  "0 = ereal r \<longleftrightarrow> r = 0"
+  unfolding zero_ereal_def by simp_all
+
+instance
+proof
+  fix a :: ereal show "0 + a = a"
+    by (cases a) (simp_all add: zero_ereal_def)
+  fix b :: ereal show "a + b = b + a"
+    by (cases rule: ereal2_cases[of a b]) simp_all
+  fix c :: ereal show "a + b + c = a + (b + c)"
+    by (cases rule: ereal3_cases[of a b c]) simp_all
+qed
+end
+
+lemma real_of_ereal_0[simp]: "real (0::ereal) = 0"
+  unfolding real_of_ereal_def zero_ereal_def by simp
+
+lemma abs_ereal_zero[simp]: "\<bar>0\<bar> = (0::ereal)"
+  unfolding zero_ereal_def abs_ereal.simps by simp
+
+lemma ereal_uminus_zero[simp]:
+  "- 0 = (0::ereal)"
+  by (simp add: zero_ereal_def)
+
+lemma ereal_uminus_zero_iff[simp]:
+  fixes a :: ereal shows "-a = 0 \<longleftrightarrow> a = 0"
+  by (cases a) simp_all
+
+lemma ereal_plus_eq_PInfty[simp]:
+  fixes a b :: ereal shows "a + b = \<infinity> \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_plus_eq_MInfty[simp]:
+  fixes a b :: ereal shows "a + b = -\<infinity> \<longleftrightarrow>
+    (a = -\<infinity> \<or> b = -\<infinity>) \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_add_cancel_left:
+  fixes a b :: ereal assumes "a \<noteq> -\<infinity>"
+  shows "a + b = a + c \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
+  using assms by (cases rule: ereal3_cases[of a b c]) auto
+
+lemma ereal_add_cancel_right:
+  fixes a b :: ereal assumes "a \<noteq> -\<infinity>"
+  shows "b + a = c + a \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
+  using assms by (cases rule: ereal3_cases[of a b c]) auto
+
+lemma ereal_real:
+  "ereal (real x) = (if \<bar>x\<bar> = \<infinity> then 0 else x)"
+  by (cases x) simp_all
+
+lemma real_of_ereal_add:
+  fixes a b :: ereal
+  shows "real (a + b) = (if (\<bar>a\<bar> = \<infinity>) \<and> (\<bar>b\<bar> = \<infinity>) \<or> (\<bar>a\<bar> \<noteq> \<infinity>) \<and> (\<bar>b\<bar> \<noteq> \<infinity>) then real a + real b else 0)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+subsubsection "Linear order on @{typ ereal}"
+
+instantiation ereal :: linorder
+begin
+
+function less_ereal where
+"   ereal x < ereal y     \<longleftrightarrow> x < y" |
+"(\<infinity>::ereal) < a           \<longleftrightarrow> False" |
+"         a < -(\<infinity>::ereal) \<longleftrightarrow> False" |
+"ereal x    < \<infinity>           \<longleftrightarrow> True" |
+"        -\<infinity> < ereal r     \<longleftrightarrow> True" |
+"        -\<infinity> < (\<infinity>::ereal) \<longleftrightarrow> True"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a,b)" by (cases x) auto
+  ultimately show P by (cases rule: ereal2_cases[of a b]) auto
+qed simp_all
+termination by (relation "{}") simp
+
+definition "x \<le> (y::ereal) \<longleftrightarrow> x < y \<or> x = y"
+
+lemma ereal_infty_less[simp]:
+  fixes x :: ereal
+  shows "x < \<infinity> \<longleftrightarrow> (x \<noteq> \<infinity>)"
+    "-\<infinity> < x \<longleftrightarrow> (x \<noteq> -\<infinity>)"
+  by (cases x, simp_all) (cases x, simp_all)
+
+lemma ereal_infty_less_eq[simp]:
+  fixes x :: ereal
+  shows "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
+  "x \<le> -\<infinity> \<longleftrightarrow> x = -\<infinity>"
+  by (auto simp add: less_eq_ereal_def)
+
+lemma ereal_less[simp]:
+  "ereal r < 0 \<longleftrightarrow> (r < 0)"
+  "0 < ereal r \<longleftrightarrow> (0 < r)"
+  "0 < (\<infinity>::ereal)"
+  "-(\<infinity>::ereal) < 0"
+  by (simp_all add: zero_ereal_def)
+
+lemma ereal_less_eq[simp]:
+  "x \<le> (\<infinity>::ereal)"
+  "-(\<infinity>::ereal) \<le> x"
+  "ereal r \<le> ereal p \<longleftrightarrow> r \<le> p"
+  "ereal r \<le> 0 \<longleftrightarrow> r \<le> 0"
+  "0 \<le> ereal r \<longleftrightarrow> 0 \<le> r"
+  by (auto simp add: less_eq_ereal_def zero_ereal_def)
+
+lemma ereal_infty_less_eq2:
+  "a \<le> b \<Longrightarrow> a = \<infinity> \<Longrightarrow> b = (\<infinity>::ereal)"
+  "a \<le> b \<Longrightarrow> b = -\<infinity> \<Longrightarrow> a = -(\<infinity>::ereal)"
+  by simp_all
+
+instance
+proof
+  fix x :: ereal show "x \<le> x"
+    by (cases x) simp_all
+  fix y :: ereal show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
+    by (cases rule: ereal2_cases[of x y]) auto
+  show "x \<le> y \<or> y \<le> x "
+    by (cases rule: ereal2_cases[of x y]) auto
+  { assume "x \<le> y" "y \<le> x" then show "x = y"
+    by (cases rule: ereal2_cases[of x y]) auto }
+  { fix z assume "x \<le> y" "y \<le> z" then show "x \<le> z"
+    by (cases rule: ereal3_cases[of x y z]) auto }
+qed
+end
+
+instance ereal :: ordered_ab_semigroup_add
+proof
+  fix a b c :: ereal assume "a \<le> b" then show "c + a \<le> c + b"
+    by (cases rule: ereal3_cases[of a b c]) auto
+qed
+
+lemma real_of_ereal_positive_mono:
+  fixes x y :: ereal shows "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
+  by (cases rule: ereal2_cases[of x y]) auto
+
+lemma ereal_MInfty_lessI[intro, simp]:
+  fixes a :: ereal shows "a \<noteq> -\<infinity> \<Longrightarrow> -\<infinity> < a"
+  by (cases a) auto
+
+lemma ereal_less_PInfty[intro, simp]:
+  fixes a :: ereal shows "a \<noteq> \<infinity> \<Longrightarrow> a < \<infinity>"
+  by (cases a) auto
+
+lemma ereal_less_ereal_Ex:
+  fixes a b :: ereal
+  shows "x < ereal r \<longleftrightarrow> x = -\<infinity> \<or> (\<exists>p. p < r \<and> x = ereal p)"
+  by (cases x) auto
+
+lemma less_PInf_Ex_of_nat: "x \<noteq> \<infinity> \<longleftrightarrow> (\<exists>n::nat. x < ereal (real n))"
+proof (cases x)
+  case (real r) then show ?thesis
+    using reals_Archimedean2[of r] by simp
+qed simp_all
+
+lemma ereal_add_mono:
+  fixes a b c d :: ereal assumes "a \<le> b" "c \<le> d" shows "a + c \<le> b + d"
+  using assms
+  apply (cases a)
+  apply (cases rule: ereal3_cases[of b c d], auto)
+  apply (cases rule: ereal3_cases[of b c d], auto)
+  done
+
+lemma ereal_minus_le_minus[simp]:
+  fixes a b :: ereal shows "- a \<le> - b \<longleftrightarrow> b \<le> a"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_minus_less_minus[simp]:
+  fixes a b :: ereal shows "- a < - b \<longleftrightarrow> b < a"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_le_real_iff:
+  "x \<le> real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> ereal x \<le> y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x \<le> 0))"
+  by (cases y) auto
+
+lemma real_le_ereal_iff:
+  "real y \<le> x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y \<le> ereal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 \<le> x))"
+  by (cases y) auto
+
+lemma ereal_less_real_iff:
+  "x < real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> ereal x < y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x < 0))"
+  by (cases y) auto
+
+lemma real_less_ereal_iff:
+  "real y < x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y < ereal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 < x))"
+  by (cases y) auto
+
+lemma real_of_ereal_pos:
+  fixes x :: ereal shows "0 \<le> x \<Longrightarrow> 0 \<le> real x" by (cases x) auto
+
+lemmas real_of_ereal_ord_simps =
+  ereal_le_real_iff real_le_ereal_iff ereal_less_real_iff real_less_ereal_iff
+
+lemma abs_ereal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: ereal\<bar> = x"
+  by (cases x) auto
+
+lemma abs_ereal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: ereal\<bar> = -x"
+  by (cases x) auto
+
+lemma abs_ereal_pos[simp]: "0 \<le> \<bar>x :: ereal\<bar>"
+  by (cases x) auto
+
+lemma real_of_ereal_le_0[simp]: "real (x :: ereal) \<le> 0 \<longleftrightarrow> (x \<le> 0 \<or> x = \<infinity>)"
+  by (cases x) auto
+
+lemma abs_real_of_ereal[simp]: "\<bar>real (x :: ereal)\<bar> = real \<bar>x\<bar>"
+  by (cases x) auto
+
+lemma zero_less_real_of_ereal:
+  fixes x :: ereal shows "0 < real x \<longleftrightarrow> (0 < x \<and> x \<noteq> \<infinity>)"
+  by (cases x) auto
+
+lemma ereal_0_le_uminus_iff[simp]:
+  fixes a :: ereal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
+  by (cases rule: ereal2_cases[of a]) auto
+
+lemma ereal_uminus_le_0_iff[simp]:
+  fixes a :: ereal shows "-a \<le> 0 \<longleftrightarrow> 0 \<le> a"
+  by (cases rule: ereal2_cases[of a]) auto
+
+lemma ereal_dense2: "x < y \<Longrightarrow> \<exists>z. x < ereal z \<and> ereal z < y"
+  using lt_ex gt_ex dense by (cases x y rule: ereal2_cases) auto
+
+lemma ereal_dense:
+  fixes x y :: ereal assumes "x < y"
+  shows "\<exists>z. x < z \<and> z < y"
+  using ereal_dense2[OF `x < y`] by blast
+
+lemma ereal_add_strict_mono:
+  fixes a b c d :: ereal
+  assumes "a = b" "0 \<le> a" "a \<noteq> \<infinity>" "c < d"
+  shows "a + c < b + d"
+  using assms by (cases rule: ereal3_cases[case_product ereal_cases, of a b c d]) auto
+
+lemma ereal_less_add: 
+  fixes a b c :: ereal shows "\<bar>a\<bar> \<noteq> \<infinity> \<Longrightarrow> c < b \<Longrightarrow> a + c < a + b"
+  by (cases rule: ereal2_cases[of b c]) auto
+
+lemma ereal_uminus_eq_reorder: "- a = b \<longleftrightarrow> a = (-b::ereal)" by auto
+
+lemma ereal_uminus_less_reorder: "- a < b \<longleftrightarrow> -b < (a::ereal)"
+  by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_less_minus)
+
+lemma ereal_uminus_le_reorder: "- a \<le> b \<longleftrightarrow> -b \<le> (a::ereal)"
+  by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_le_minus)
+
+lemmas ereal_uminus_reorder =
+  ereal_uminus_eq_reorder ereal_uminus_less_reorder ereal_uminus_le_reorder
+
+lemma ereal_bot:
+  fixes x :: ereal assumes "\<And>B. x \<le> ereal B" shows "x = - \<infinity>"
+proof (cases x)
+  case (real r) with assms[of "r - 1"] show ?thesis by auto
+next case PInf with assms[of 0] show ?thesis by auto
+next case MInf then show ?thesis by simp
+qed
+
+lemma ereal_top:
+  fixes x :: ereal assumes "\<And>B. x \<ge> ereal B" shows "x = \<infinity>"
+proof (cases x)
+  case (real r) with assms[of "r + 1"] show ?thesis by auto
+next case MInf with assms[of 0] show ?thesis by auto
+next case PInf then show ?thesis by simp
+qed
+
+lemma
+  shows ereal_max[simp]: "ereal (max x y) = max (ereal x) (ereal y)"
+    and ereal_min[simp]: "ereal (min x y) = min (ereal x) (ereal y)"
+  by (simp_all add: min_def max_def)
+
+lemma ereal_max_0: "max 0 (ereal r) = ereal (max 0 r)"
+  by (auto simp: zero_ereal_def)
+
+lemma
+  fixes f :: "nat \<Rightarrow> ereal"
+  shows incseq_uminus[simp]: "incseq (\<lambda>x. - f x) \<longleftrightarrow> decseq f"
+  and decseq_uminus[simp]: "decseq (\<lambda>x. - f x) \<longleftrightarrow> incseq f"
+  unfolding decseq_def incseq_def by auto
+
+lemma incseq_ereal: "incseq f \<Longrightarrow> incseq (\<lambda>x. ereal (f x))"
+  unfolding incseq_def by auto
+
+lemma ereal_add_nonneg_nonneg:
+  fixes a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
+  using add_mono[of 0 a 0 b] by simp
+
+lemma image_eqD: "f ` A = B \<Longrightarrow> (\<forall>x\<in>A. f x \<in> B)"
+  by auto
+
+lemma incseq_setsumI:
+  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
+  assumes "\<And>i. 0 \<le> f i"
+  shows "incseq (\<lambda>i. setsum f {..< i})"
+proof (intro incseq_SucI)
+  fix n have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
+    using assms by (rule add_left_mono)
+  then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
+    by auto
+qed
+
+lemma incseq_setsumI2:
+  fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
+  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
+  shows "incseq (\<lambda>i. \<Sum>n\<in>A. f n i)"
+  using assms unfolding incseq_def by (auto intro: setsum_mono)
+
+subsubsection "Multiplication"
+
+instantiation ereal :: "{comm_monoid_mult, sgn}"
+begin
+
+definition "1 = ereal 1"
+
+function sgn_ereal where
+  "sgn (ereal r) = ereal (sgn r)"
+| "sgn (\<infinity>::ereal) = 1"
+| "sgn (-\<infinity>::ereal) = -1"
+by (auto intro: ereal_cases)
+termination proof qed (rule wf_empty)
+
+function times_ereal where
+"ereal r * ereal p = ereal (r * p)" |
+"ereal r * \<infinity> = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
+"\<infinity> * ereal r = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
+"ereal r * -\<infinity> = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
+"-\<infinity> * ereal r = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
+"(\<infinity>::ereal) * \<infinity> = \<infinity>" |
+"-(\<infinity>::ereal) * \<infinity> = -\<infinity>" |
+"(\<infinity>::ereal) * -\<infinity> = -\<infinity>" |
+"-(\<infinity>::ereal) * -\<infinity> = \<infinity>"
+proof -
+  case (goal1 P x)
+  moreover then obtain a b where "x = (a, b)" by (cases x) auto
+  ultimately show P by (cases rule: ereal2_cases[of a b]) auto
+qed simp_all
+termination by (relation "{}") simp
+
+instance
+proof
+  fix a :: ereal show "1 * a = a"
+    by (cases a) (simp_all add: one_ereal_def)
+  fix b :: ereal show "a * b = b * a"
+    by (cases rule: ereal2_cases[of a b]) simp_all
+  fix c :: ereal show "a * b * c = a * (b * c)"
+    by (cases rule: ereal3_cases[of a b c])
+       (simp_all add: zero_ereal_def zero_less_mult_iff)
+qed
+end
+
+lemma real_of_ereal_le_1:
+  fixes a :: ereal shows "a \<le> 1 \<Longrightarrow> real a \<le> 1"
+  by (cases a) (auto simp: one_ereal_def)
+
+lemma abs_ereal_one[simp]: "\<bar>1\<bar> = (1::ereal)"
+  unfolding one_ereal_def by simp
+
+lemma ereal_mult_zero[simp]:
+  fixes a :: ereal shows "a * 0 = 0"
+  by (cases a) (simp_all add: zero_ereal_def)
+
+lemma ereal_zero_mult[simp]:
+  fixes a :: ereal shows "0 * a = 0"
+  by (cases a) (simp_all add: zero_ereal_def)
+
+lemma ereal_m1_less_0[simp]:
+  "-(1::ereal) < 0"
+  by (simp add: zero_ereal_def one_ereal_def)
+
+lemma ereal_zero_m1[simp]:
+  "1 \<noteq> (0::ereal)"
+  by (simp add: zero_ereal_def one_ereal_def)
+
+lemma ereal_times_0[simp]:
+  fixes x :: ereal shows "0 * x = 0"
+  by (cases x) (auto simp: zero_ereal_def)
+
+lemma ereal_times[simp]:
+  "1 \<noteq> (\<infinity>::ereal)" "(\<infinity>::ereal) \<noteq> 1"
+  "1 \<noteq> -(\<infinity>::ereal)" "-(\<infinity>::ereal) \<noteq> 1"
+  by (auto simp add: times_ereal_def one_ereal_def)
+
+lemma ereal_plus_1[simp]:
+  "1 + ereal r = ereal (r + 1)" "ereal r + 1 = ereal (r + 1)"
+  "1 + -(\<infinity>::ereal) = -\<infinity>" "-(\<infinity>::ereal) + 1 = -\<infinity>"
+  unfolding one_ereal_def by auto
+
+lemma ereal_zero_times[simp]:
+  fixes a b :: ereal shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_mult_eq_PInfty[simp]:
+  shows "a * b = (\<infinity>::ereal) \<longleftrightarrow>
+    (a = \<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = -\<infinity>)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_mult_eq_MInfty[simp]:
+  shows "a * b = -(\<infinity>::ereal) \<longleftrightarrow>
+    (a = \<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = -\<infinity>)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_0_less_1[simp]: "0 < (1::ereal)"
+  by (simp_all add: zero_ereal_def one_ereal_def)
+
+lemma ereal_zero_one[simp]: "0 \<noteq> (1::ereal)"
+  by (simp_all add: zero_ereal_def one_ereal_def)
+
+lemma ereal_mult_minus_left[simp]:
+  fixes a b :: ereal shows "-a * b = - (a * b)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_mult_minus_right[simp]:
+  fixes a b :: ereal shows "a * -b = - (a * b)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_mult_infty[simp]:
+  "a * (\<infinity>::ereal) = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
+  by (cases a) auto
+
+lemma ereal_infty_mult[simp]:
+  "(\<infinity>::ereal) * a = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
+  by (cases a) auto
+
+lemma ereal_mult_strict_right_mono:
+  assumes "a < b" and "0 < c" "c < (\<infinity>::ereal)"
+  shows "a * c < b * c"
+  using assms
+  by (cases rule: ereal3_cases[of a b c])
+     (auto simp: zero_le_mult_iff ereal_less_PInfty)
+
+lemma ereal_mult_strict_left_mono:
+  "\<lbrakk> a < b ; 0 < c ; c < (\<infinity>::ereal)\<rbrakk> \<Longrightarrow> c * a < c * b"
+  using ereal_mult_strict_right_mono by (simp add: mult_commute[of c])
+
+lemma ereal_mult_right_mono:
+  fixes a b c :: ereal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a*c \<le> b*c"
+  using assms
+  apply (cases "c = 0") apply simp
+  by (cases rule: ereal3_cases[of a b c])
+     (auto simp: zero_le_mult_iff ereal_less_PInfty)
+
+lemma ereal_mult_left_mono:
+  fixes a b c :: ereal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> c * a \<le> c * b"
+  using ereal_mult_right_mono by (simp add: mult_commute[of c])
+
+lemma zero_less_one_ereal[simp]: "0 \<le> (1::ereal)"
+  by (simp add: one_ereal_def zero_ereal_def)
+
+lemma ereal_0_le_mult[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * (b :: ereal)"
+  by (cases rule: ereal2_cases[of a b]) (auto simp: mult_nonneg_nonneg)
+
+lemma ereal_right_distrib:
+  fixes r a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> r * (a + b) = r * a + r * b"
+  by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps)
+
+lemma ereal_left_distrib:
+  fixes r a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> (a + b) * r = a * r + b * r"
+  by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps)
+
+lemma ereal_mult_le_0_iff:
+  fixes a b :: ereal
+  shows "a * b \<le> 0 \<longleftrightarrow> (0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b)"
+  by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_le_0_iff)
+
+lemma ereal_zero_le_0_iff:
+  fixes a b :: ereal
+  shows "0 \<le> a * b \<longleftrightarrow> (0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0)"
+  by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_le_mult_iff)
+
+lemma ereal_mult_less_0_iff:
+  fixes a b :: ereal
+  shows "a * b < 0 \<longleftrightarrow> (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b)"
+  by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_less_0_iff)
+
+lemma ereal_zero_less_0_iff:
+  fixes a b :: ereal
+  shows "0 < a * b \<longleftrightarrow> (0 < a \<and> 0 < b) \<or> (a < 0 \<and> b < 0)"
+  by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_less_mult_iff)
+
+lemma ereal_distrib:
+  fixes a b c :: ereal
+  assumes "a \<noteq> \<infinity> \<or> b \<noteq> -\<infinity>" "a \<noteq> -\<infinity> \<or> b \<noteq> \<infinity>" "\<bar>c\<bar> \<noteq> \<infinity>"
+  shows "(a + b) * c = a * c + b * c"
+  using assms
+  by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps)
+
+lemma ereal_le_epsilon:
+  fixes x y :: ereal
+  assumes "ALL e. 0 < e --> x <= y + e"
+  shows "x <= y"
+proof-
+{ assume a: "EX r. y = ereal r"
+  from this obtain r where r_def: "y = ereal r" by auto
+  { assume "x=(-\<infinity>)" hence ?thesis by auto }
+  moreover
+  { assume "~(x=(-\<infinity>))"
+    from this obtain p where p_def: "x = ereal p"
+    using a assms[rule_format, of 1] by (cases x) auto
+    { fix e have "0 < e --> p <= r + e"
+      using assms[rule_format, of "ereal e"] p_def r_def by auto }
+    hence "p <= r" apply (subst field_le_epsilon) by auto
+    hence ?thesis using r_def p_def by auto
+  } ultimately have ?thesis by blast
+}
+moreover
+{ assume "y=(-\<infinity>) | y=\<infinity>" hence ?thesis
+    using assms[rule_format, of 1] by (cases x) auto
+} ultimately show ?thesis by (cases y) auto
+qed
+
+
+lemma ereal_le_epsilon2:
+  fixes x y :: ereal
+  assumes "ALL e. 0 < e --> x <= y + ereal e"
+  shows "x <= y"
+proof-
+{ fix e :: ereal assume "e>0"
+  { assume "e=\<infinity>" hence "x<=y+e" by auto }
+  moreover
+  { assume "e~=\<infinity>"
+    from this obtain r where "e = ereal r" using `e>0` apply (cases e) by auto
+    hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto
+  } ultimately have "x<=y+e" by blast
+} from this show ?thesis using ereal_le_epsilon by auto
+qed
+
+lemma ereal_le_real:
+  fixes x y :: ereal
+  assumes "ALL z. x <= ereal z --> y <= ereal z"
+  shows "y <= x"
+by (metis assms ereal_bot ereal_cases ereal_infty_less_eq ereal_less_eq linorder_le_cases)
+
+lemma ereal_le_ereal:
+  fixes x y :: ereal
+  assumes "\<And>B. B < x \<Longrightarrow> B <= y"
+  shows "x <= y"
+by (metis assms ereal_dense leD linorder_le_less_linear)
+
+lemma ereal_ge_ereal:
+  fixes x y :: ereal
+  assumes "ALL B. B>x --> B >= y"
+  shows "x >= y"
+by (metis assms ereal_dense leD linorder_le_less_linear)
+
+lemma setprod_ereal_0:
+  fixes f :: "'a \<Rightarrow> ereal"
+  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
+proof cases
+  assume "finite A"
+  then show ?thesis by (induct A) auto
+qed auto
+
+lemma setprod_ereal_pos:
+  fixes f :: "'a \<Rightarrow> ereal" assumes pos: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" shows "0 \<le> (\<Prod>i\<in>I. f i)"
+proof cases
+  assume "finite I" from this pos show ?thesis by induct auto
+qed simp
+
+lemma setprod_PInf:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
+  shows "(\<Prod>i\<in>I. f i) = \<infinity> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
+proof cases
+  assume "finite I" from this assms show ?thesis
+  proof (induct I)
+    case (insert i I)
+    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_ereal_pos)
+    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
+    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
+      using setprod_ereal_pos[of I f] pos
+      by (cases rule: ereal2_cases[of "f i" "setprod f I"]) auto
+    also have "\<dots> \<longleftrightarrow> finite (insert i I) \<and> (\<exists>j\<in>insert i I. f j = \<infinity>) \<and> (\<forall>j\<in>insert i I. f j \<noteq> 0)"
+      using insert by (auto simp: setprod_ereal_0)
+    finally show ?case .
+  qed simp
+qed simp
+
+lemma setprod_ereal: "(\<Prod>i\<in>A. ereal (f i)) = ereal (setprod f A)"
+proof cases
+  assume "finite A" then show ?thesis
+    by induct (auto simp: one_ereal_def)
+qed (simp add: one_ereal_def)
+
+subsubsection {* Power *}
+
+instantiation ereal :: power
+begin
+primrec power_ereal where
+  "power_ereal x 0 = 1" |
+  "power_ereal x (Suc n) = x * x ^ n"
+instance ..
+end
+
+lemma ereal_power[simp]: "(ereal x) ^ n = ereal (x^n)"
+  by (induct n) (auto simp: one_ereal_def)
+
+lemma ereal_power_PInf[simp]: "(\<infinity>::ereal) ^ n = (if n = 0 then 1 else \<infinity>)"
+  by (induct n) (auto simp: one_ereal_def)
+
+lemma ereal_power_uminus[simp]:
+  fixes x :: ereal
+  shows "(- x) ^ n = (if even n then x ^ n else - (x^n))"
+  by (induct n) (auto simp: one_ereal_def)
+
+lemma ereal_power_number_of[simp]:
+  "(number_of num :: ereal) ^ n = ereal (number_of num ^ n)"
+  by (induct n) (auto simp: one_ereal_def)
+
+lemma zero_le_power_ereal[simp]:
+  fixes a :: ereal assumes "0 \<le> a"
+  shows "0 \<le> a ^ n"
+  using assms by (induct n) (auto simp: ereal_zero_le_0_iff)
+
+subsubsection {* Subtraction *}
+
+lemma ereal_minus_minus_image[simp]:
+  fixes S :: "ereal set"
+  shows "uminus ` uminus ` S = S"
+  by (auto simp: image_iff)
+
+lemma ereal_uminus_lessThan[simp]:
+  fixes a :: ereal shows "uminus ` {..<a} = {-a<..}"
+proof (safe intro!: image_eqI)
+  fix x assume "-a < x"
+  then have "- x < - (- a)" by (simp del: ereal_uminus_uminus)
+  then show "- x < a" by simp
+qed auto
+
+lemma ereal_uminus_greaterThan[simp]:
+  "uminus ` {(a::ereal)<..} = {..<-a}"
+  by (metis ereal_uminus_lessThan ereal_uminus_uminus
+            ereal_minus_minus_image)
+
+instantiation ereal :: minus
+begin
+definition "x - y = x + -(y::ereal)"
+instance ..
+end
+
+lemma ereal_minus[simp]:
+  "ereal r - ereal p = ereal (r - p)"
+  "-\<infinity> - ereal r = -\<infinity>"
+  "ereal r - \<infinity> = -\<infinity>"
+  "(\<infinity>::ereal) - x = \<infinity>"
+  "-(\<infinity>::ereal) - \<infinity> = -\<infinity>"
+  "x - -y = x + y"
+  "x - 0 = x"
+  "0 - x = -x"
+  by (simp_all add: minus_ereal_def)
+
+lemma ereal_x_minus_x[simp]:
+  "x - x = (if \<bar>x\<bar> = \<infinity> then \<infinity> else 0::ereal)"
+  by (cases x) simp_all
+
+lemma ereal_eq_minus_iff:
+  fixes x y z :: ereal
+  shows "x = z - y \<longleftrightarrow>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y = z) \<and>
+    (y = -\<infinity> \<longrightarrow> x = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> z = \<infinity> \<longrightarrow> x = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>)"
+  by (cases rule: ereal3_cases[of x y z]) auto
+
+lemma ereal_eq_minus:
+  fixes x y z :: ereal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x = z - y \<longleftrightarrow> x + y = z"
+  by (auto simp: ereal_eq_minus_iff)
+
+lemma ereal_less_minus_iff:
+  fixes x y z :: ereal
+  shows "x < z - y \<longleftrightarrow>
+    (y = \<infinity> \<longrightarrow> z = \<infinity> \<and> x \<noteq> \<infinity>) \<and>
+    (y = -\<infinity> \<longrightarrow> x \<noteq> \<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity>\<longrightarrow> x + y < z)"
+  by (cases rule: ereal3_cases[of x y z]) auto
+
+lemma ereal_less_minus:
+  fixes x y z :: ereal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x < z - y \<longleftrightarrow> x + y < z"
+  by (auto simp: ereal_less_minus_iff)
+
+lemma ereal_le_minus_iff:
+  fixes x y z :: ereal
+  shows "x \<le> z - y \<longleftrightarrow>
+    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y \<le> z)"
+  by (cases rule: ereal3_cases[of x y z]) auto
+
+lemma ereal_le_minus:
+  fixes x y z :: ereal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x \<le> z - y \<longleftrightarrow> x + y \<le> z"
+  by (auto simp: ereal_le_minus_iff)
+
+lemma ereal_minus_less_iff:
+  fixes x y z :: ereal
+  shows "x - y < z \<longleftrightarrow>
+    y \<noteq> -\<infinity> \<and> (y = \<infinity> \<longrightarrow> x \<noteq> \<infinity> \<and> z \<noteq> -\<infinity>) \<and>
+    (y \<noteq> \<infinity> \<longrightarrow> x < z + y)"
+  by (cases rule: ereal3_cases[of x y z]) auto
+
+lemma ereal_minus_less:
+  fixes x y z :: ereal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y < z \<longleftrightarrow> x < z + y"
+  by (auto simp: ereal_minus_less_iff)
+
+lemma ereal_minus_le_iff:
+  fixes x y z :: ereal
+  shows "x - y \<le> z \<longleftrightarrow>
+    (y = -\<infinity> \<longrightarrow> z = \<infinity>) \<and>
+    (y = \<infinity> \<longrightarrow> x = \<infinity> \<longrightarrow> z = \<infinity>) \<and>
+    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x \<le> z + y)"
+  by (cases rule: ereal3_cases[of x y z]) auto
+
+lemma ereal_minus_le:
+  fixes x y z :: ereal
+  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y \<le> z \<longleftrightarrow> x \<le> z + y"
+  by (auto simp: ereal_minus_le_iff)
+
+lemma ereal_minus_eq_minus_iff:
+  fixes a b c :: ereal
+  shows "a - b = a - c \<longleftrightarrow>
+    b = c \<or> a = \<infinity> \<or> (a = -\<infinity> \<and> b \<noteq> -\<infinity> \<and> c \<noteq> -\<infinity>)"
+  by (cases rule: ereal3_cases[of a b c]) auto
+
+lemma ereal_add_le_add_iff:
+  fixes a b c :: ereal
+  shows "c + a \<le> c + b \<longleftrightarrow>
+    a \<le> b \<or> c = \<infinity> \<or> (c = -\<infinity> \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>)"
+  by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps)
+
+lemma ereal_mult_le_mult_iff:
+  fixes a b c :: ereal
+  shows "\<bar>c\<bar> \<noteq> \<infinity> \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
+  by (cases rule: ereal3_cases[of a b c]) (simp_all add: mult_le_cancel_left)
+
+lemma ereal_minus_mono:
+  fixes A B C D :: ereal assumes "A \<le> B" "D \<le> C"
+  shows "A - C \<le> B - D"
+  using assms
+  by (cases rule: ereal3_cases[case_product ereal_cases, of A B C D]) simp_all
+
+lemma real_of_ereal_minus:
+  fixes a b :: ereal
+  shows "real (a - b) = (if \<bar>a\<bar> = \<infinity> \<or> \<bar>b\<bar> = \<infinity> then 0 else real a - real b)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_diff_positive:
+  fixes a b :: ereal shows "a \<le> b \<Longrightarrow> 0 \<le> b - a"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_between:
+  fixes x e :: ereal
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>" "0 < e"
+  shows "x - e < x" "x < x + e"
+using assms apply (cases x, cases e) apply auto
+using assms by (cases x, cases e) auto
+
+subsubsection {* Division *}
+
+instantiation ereal :: inverse
+begin
+
+function inverse_ereal where
+"inverse (ereal r) = (if r = 0 then \<infinity> else ereal (inverse r))" |
+"inverse (\<infinity>::ereal) = 0" |
+"inverse (-\<infinity>::ereal) = 0"
+  by (auto intro: ereal_cases)
+termination by (relation "{}") simp
+
+definition "x / y = x * inverse (y :: ereal)"
+
+instance proof qed
+end
+
+lemma real_of_ereal_inverse[simp]:
+  fixes a :: ereal
+  shows "real (inverse a) = 1 / real a"
+  by (cases a) (auto simp: inverse_eq_divide)
+
+lemma ereal_inverse[simp]:
+  "inverse (0::ereal) = \<infinity>"
+  "inverse (1::ereal) = 1"
+  by (simp_all add: one_ereal_def zero_ereal_def)
+
+lemma ereal_divide[simp]:
+  "ereal r / ereal p = (if p = 0 then ereal r * \<infinity> else ereal (r / p))"
+  unfolding divide_ereal_def by (auto simp: divide_real_def)
+
+lemma ereal_divide_same[simp]:
+  fixes x :: ereal shows "x / x = (if \<bar>x\<bar> = \<infinity> \<or> x = 0 then 0 else 1)"
+  by (cases x)
+     (simp_all add: divide_real_def divide_ereal_def one_ereal_def)
+
+lemma ereal_inv_inv[simp]:
+  fixes x :: ereal shows "inverse (inverse x) = (if x \<noteq> -\<infinity> then x else \<infinity>)"
+  by (cases x) auto
+
+lemma ereal_inverse_minus[simp]:
+  fixes x :: ereal shows "inverse (- x) = (if x = 0 then \<infinity> else -inverse x)"
+  by (cases x) simp_all
+
+lemma ereal_uminus_divide[simp]:
+  fixes x y :: ereal shows "- x / y = - (x / y)"
+  unfolding divide_ereal_def by simp
+
+lemma ereal_divide_Infty[simp]:
+  fixes x :: ereal shows "x / \<infinity> = 0" "x / -\<infinity> = 0"
+  unfolding divide_ereal_def by simp_all
+
+lemma ereal_divide_one[simp]:
+  "x / 1 = (x::ereal)"
+  unfolding divide_ereal_def by simp
+
+lemma ereal_divide_ereal[simp]:
+  "\<infinity> / ereal r = (if 0 \<le> r then \<infinity> else -\<infinity>)"
+  unfolding divide_ereal_def by simp
+
+lemma zero_le_divide_ereal[simp]:
+  fixes a :: ereal assumes "0 \<le> a" "0 \<le> b"
+  shows "0 \<le> a / b"
+  using assms by (cases rule: ereal2_cases[of a b]) (auto simp: zero_le_divide_iff)
+
+lemma ereal_le_divide_pos:
+  fixes x y z :: ereal shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> x * y \<le> z"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_divide_le_pos:
+  fixes x y z :: ereal shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> z \<le> x * y"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_le_divide_neg:
+  fixes x y z :: ereal shows "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> z \<le> x * y"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_divide_le_neg:
+  fixes x y z :: ereal shows "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> x * y \<le> z"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_inverse_antimono_strict:
+  fixes x y :: ereal
+  shows "0 \<le> x \<Longrightarrow> x < y \<Longrightarrow> inverse y < inverse x"
+  by (cases rule: ereal2_cases[of x y]) auto
+
+lemma ereal_inverse_antimono:
+  fixes x y :: ereal
+  shows "0 \<le> x \<Longrightarrow> x <= y \<Longrightarrow> inverse y <= inverse x"
+  by (cases rule: ereal2_cases[of x y]) auto
+
+lemma inverse_inverse_Pinfty_iff[simp]:
+  fixes x :: ereal shows "inverse x = \<infinity> \<longleftrightarrow> x = 0"
+  by (cases x) auto
+
+lemma ereal_inverse_eq_0:
+  fixes x :: ereal shows "inverse x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity>"
+  by (cases x) auto
+
+lemma ereal_0_gt_inverse:
+  fixes x :: ereal shows "0 < inverse x \<longleftrightarrow> x \<noteq> \<infinity> \<and> 0 \<le> x"
+  by (cases x) auto
+
+lemma ereal_mult_less_right:
+  fixes a b c :: ereal
+  assumes "b * a < c * a" "0 < a" "a < \<infinity>"
+  shows "b < c"
+  using assms
+  by (cases rule: ereal3_cases[of a b c])
+     (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
+
+lemma ereal_power_divide:
+  fixes x y :: ereal shows "y \<noteq> 0 \<Longrightarrow> (x / y) ^ n = x^n / y^n"
+  by (cases rule: ereal2_cases[of x y])
+     (auto simp: one_ereal_def zero_ereal_def power_divide not_le
+                 power_less_zero_eq zero_le_power_iff)
+
+lemma ereal_le_mult_one_interval:
+  fixes x y :: ereal
+  assumes y: "y \<noteq> -\<infinity>"
+  assumes z: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
+  shows "x \<le> y"
+proof (cases x)
+  case PInf with z[of "1 / 2"] show "x \<le> y" by (simp add: one_ereal_def)
+next
+  case (real r) note r = this
+  show "x \<le> y"
+  proof (cases y)
+    case (real p) note p = this
+    have "r \<le> p"
+    proof (rule field_le_mult_one_interval)
+      fix z :: real assume "0 < z" and "z < 1"
+      with z[of "ereal z"]
+      show "z * r \<le> p" using p r by (auto simp: zero_le_mult_iff one_ereal_def)
+    qed
+    then show "x \<le> y" using p r by simp
+  qed (insert y, simp_all)
+qed simp
+
+subsection "Complete lattice"
+
+instantiation ereal :: lattice
+begin
+definition [simp]: "sup x y = (max x y :: ereal)"
+definition [simp]: "inf x y = (min x y :: ereal)"
+instance proof qed simp_all
+end
+
+instantiation ereal :: complete_lattice
+begin
+
+definition "bot = (-\<infinity>::ereal)"
+definition "top = (\<infinity>::ereal)"
+
+definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: ereal)"
+definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: ereal)"
+
+lemma ereal_complete_Sup:
+  fixes S :: "ereal set" assumes "S \<noteq> {}"
+  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
+proof cases
+  assume "\<exists>x. \<forall>a\<in>S. a \<le> ereal x"
+  then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> ereal y" by auto
+  then have "\<infinity> \<notin> S" by force
+  show ?thesis
+  proof cases
+    assume "S = {-\<infinity>}"
+    then show ?thesis by (auto intro!: exI[of _ "-\<infinity>"])
+  next
+    assume "S \<noteq> {-\<infinity>}"
+    with `S \<noteq> {}` `\<infinity> \<notin> S` obtain x where "x \<in> S - {-\<infinity>}" "x \<noteq> \<infinity>" by auto
+    with y `\<infinity> \<notin> S` have "\<forall>z\<in>real ` (S - {-\<infinity>}). z \<le> y"
+      by (auto simp: real_of_ereal_ord_simps)
+    with reals_complete2[of "real ` (S - {-\<infinity>})"] `x \<in> S - {-\<infinity>}`
+    obtain s where s:
+       "\<forall>y\<in>S - {-\<infinity>}. real y \<le> s" "\<And>z. (\<forall>y\<in>S - {-\<infinity>}. real y \<le> z) \<Longrightarrow> s \<le> z"
+       by auto
+    show ?thesis
+    proof (safe intro!: exI[of _ "ereal s"])
+      fix z assume "z \<in> S" with `\<infinity> \<notin> S` show "z \<le> ereal s"
+      proof (cases z)
+        case (real r)
+        then show ?thesis
+          using s(1)[rule_format, of z] `z \<in> S` `z = ereal r` by auto
+      qed auto
+    next
+      fix z assume *: "\<forall>y\<in>S. y \<le> z"
+      with `S \<noteq> {-\<infinity>}` `S \<noteq> {}` show "ereal s \<le> z"
+      proof (cases z)
+        case (real u)
+        with * have "s \<le> u"
+          by (intro s(2)[of u]) (auto simp: real_of_ereal_ord_simps)
+        then show ?thesis using real by simp
+      qed auto
+    qed
+  qed
+next
+  assume *: "\<not> (\<exists>x. \<forall>a\<in>S. a \<le> ereal x)"
+  show ?thesis
+  proof (safe intro!: exI[of _ \<infinity>])
+    fix y assume **: "\<forall>z\<in>S. z \<le> y"
+    with * show "\<infinity> \<le> y"
+    proof (cases y)
+      case MInf with * ** show ?thesis by (force simp: not_le)
+    qed auto
+  qed simp
+qed
+
+lemma ereal_complete_Inf:
+  fixes S :: "ereal set" assumes "S ~= {}"
+  shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)"
+proof-
+def S1 == "uminus ` S"
+hence "S1 ~= {}" using assms by auto
+from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)"
+   using ereal_complete_Sup[of S1] by auto
+{ fix z assume "ALL y:S. z <= y"
+  hence "ALL y:S1. y <= -z" unfolding S1_def by auto
+  hence "x <= -z" using x_def by auto
+  hence "z <= -x"
+    apply (subst ereal_uminus_uminus[symmetric])
+    unfolding ereal_minus_le_minus . }
+moreover have "(ALL y:S. -x <= y)"
+   using x_def unfolding S1_def
+   apply simp
+   apply (subst (3) ereal_uminus_uminus[symmetric])
+   unfolding ereal_minus_le_minus by simp
+ultimately show ?thesis by auto
+qed
+
+lemma ereal_complete_uminus_eq:
+  fixes S :: "ereal set"
+  shows "(\<forall>y\<in>uminus`S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>uminus`S. y \<le> z) \<longrightarrow> x \<le> z)
+     \<longleftrightarrow> (\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
+  by simp (metis ereal_minus_le_minus ereal_uminus_uminus)
+
+lemma ereal_Sup_uminus_image_eq:
+  fixes S :: "ereal set"
+  shows "Sup (uminus ` S) = - Inf S"
+proof cases
+  assume "S = {}"
+  moreover have "(THE x. All (op \<le> x)) = (-\<infinity>::ereal)"
+    by (rule the_equality) (auto intro!: ereal_bot)
+  moreover have "(SOME x. \<forall>y. y \<le> x) = (\<infinity>::ereal)"
+    by (rule some_equality) (auto intro!: ereal_top)
+  ultimately show ?thesis unfolding Inf_ereal_def Sup_ereal_def
+    Least_def Greatest_def GreatestM_def by simp
+next
+  assume "S \<noteq> {}"
+  with ereal_complete_Sup[of "uminus`S"]
+  obtain x where x: "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
+    unfolding ereal_complete_uminus_eq by auto
+  show "Sup (uminus ` S) = - Inf S"
+    unfolding Inf_ereal_def Greatest_def GreatestM_def
+  proof (intro someI2[of _ _ "\<lambda>x. Sup (uminus`S) = - x"])
+    show "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> -x)"
+      using x .
+    fix x' assume "(\<forall>y\<in>S. x' \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> x')"
+    then have "(\<forall>y\<in>uminus`S. y \<le> - x') \<and> (\<forall>y. (\<forall>z\<in>uminus`S. z \<le> y) \<longrightarrow> - x' \<le> y)"
+      unfolding ereal_complete_uminus_eq by simp
+    then show "Sup (uminus ` S) = -x'"
+      unfolding Sup_ereal_def ereal_uminus_eq_iff
+      by (intro Least_equality) auto
+  qed
+qed
+
+instance
+proof
+  { fix x :: ereal and A
+    show "bot <= x" by (cases x) (simp_all add: bot_ereal_def)
+    show "x <= top" by (simp add: top_ereal_def) }
+
+  { fix x :: ereal and A assume "x : A"
+    with ereal_complete_Sup[of A]
+    obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
+    hence "x <= s" using `x : A` by auto
+    also have "... = Sup A" using s unfolding Sup_ereal_def
+      by (auto intro!: Least_equality[symmetric])
+    finally show "x <= Sup A" . }
+  note le_Sup = this
+
+  { fix x :: ereal and A assume *: "!!z. (z : A ==> z <= x)"
+    show "Sup A <= x"
+    proof (cases "A = {}")
+      case True
+      hence "Sup A = -\<infinity>" unfolding Sup_ereal_def
+        by (auto intro!: Least_equality)
+      thus "Sup A <= x" by simp
+    next
+      case False
+      with ereal_complete_Sup[of A]
+      obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
+      hence "Sup A = s"
+        unfolding Sup_ereal_def by (auto intro!: Least_equality)
+      also have "s <= x" using * s by auto
+      finally show "Sup A <= x" .
+    qed }
+  note Sup_le = this
+
+  { fix x :: ereal and A assume "x \<in> A"
+    with le_Sup[of "-x" "uminus`A"] show "Inf A \<le> x"
+      unfolding ereal_Sup_uminus_image_eq by simp }
+
+  { fix x :: ereal and A assume *: "!!z. (z : A ==> x <= z)"
+    with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
+      unfolding ereal_Sup_uminus_image_eq by force }
+qed
+end
+
+lemma ereal_SUPR_uminus:
+  fixes f :: "'a => ereal"
+  shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
+  unfolding SUPR_def INFI_def
+  using ereal_Sup_uminus_image_eq[of "f`R"]
+  by (simp add: image_image)
+
+lemma ereal_INFI_uminus:
+  fixes f :: "'a => ereal"
+  shows "(INF i : R. -(f i)) = -(SUP i : R. f i)"
+  using ereal_SUPR_uminus[of _ "\<lambda>x. - f x"] by simp
+
+lemma ereal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::ereal set)"
+  using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image)
+
+lemma ereal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: ereal set)"
+  by (auto intro!: inj_onI)
+
+lemma ereal_image_uminus_shift:
+  fixes X Y :: "ereal set" shows "uminus ` X = Y \<longleftrightarrow> X = uminus ` Y"
+proof
+  assume "uminus ` X = Y"
+  then have "uminus ` uminus ` X = uminus ` Y"
+    by (simp add: inj_image_eq_iff)
+  then show "X = uminus ` Y" by (simp add: image_image)
+qed (simp add: image_image)
+
+lemma Inf_ereal_iff:
+  fixes z :: ereal
+  shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x<y) <-> Inf X < y"
+  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
+            order_less_le_trans)
+
+lemma Sup_eq_MInfty:
+  fixes S :: "ereal set" shows "Sup S = -\<infinity> \<longleftrightarrow> S = {} \<or> S = {-\<infinity>}"
+proof
+  assume a: "Sup S = -\<infinity>"
+  with complete_lattice_class.Sup_upper[of _ S]
+  show "S={} \<or> S={-\<infinity>}" by auto
+next
+  assume "S={} \<or> S={-\<infinity>}" then show "Sup S = -\<infinity>"
+    unfolding Sup_ereal_def by (auto intro!: Least_equality)
+qed
+
+lemma Inf_eq_PInfty:
+  fixes S :: "ereal set" shows "Inf S = \<infinity> \<longleftrightarrow> S = {} \<or> S = {\<infinity>}"
+  using Sup_eq_MInfty[of "uminus`S"]
+  unfolding ereal_Sup_uminus_image_eq ereal_image_uminus_shift by simp
+
+lemma Inf_eq_MInfty: 
+  fixes S :: "ereal set" shows "-\<infinity> \<in> S \<Longrightarrow> Inf S = -\<infinity>"
+  unfolding Inf_ereal_def
+  by (auto intro!: Greatest_equality)
+
+lemma Sup_eq_PInfty:
+  fixes S :: "ereal set" shows "\<infinity> \<in> S \<Longrightarrow> Sup S = \<infinity>"
+  unfolding Sup_ereal_def
+  by (auto intro!: Least_equality)
+
+lemma ereal_SUPI:
+  fixes x :: ereal
+  assumes "!!i. i : A ==> f i <= x"
+  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
+  shows "(SUP i:A. f i) = x"
+  unfolding SUPR_def Sup_ereal_def
+  using assms by (auto intro!: Least_equality)
+
+lemma ereal_INFI:
+  fixes x :: ereal
+  assumes "!!i. i : A ==> f i >= x"
+  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
+  shows "(INF i:A. f i) = x"
+  unfolding INFI_def Inf_ereal_def
+  using assms by (auto intro!: Greatest_equality)
+
+lemma Sup_ereal_close:
+  fixes e :: ereal
+  assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
+  shows "\<exists>x\<in>S. Sup S - e < x"
+  using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1])
+
+lemma Inf_ereal_close:
+  fixes e :: ereal assumes "\<bar>Inf X\<bar> \<noteq> \<infinity>" "0 < e"
+  shows "\<exists>x\<in>X. x < Inf X + e"
+proof (rule Inf_less_iff[THEN iffD1])
+  show "Inf X < Inf X + e" using assms
+    by (cases e) auto
+qed
+
+lemma Sup_eq_top_iff:
+  fixes A :: "'a::{complete_lattice, linorder} set"
+  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
+proof
+  assume *: "Sup A = top"
+  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
+  proof (intro allI impI)
+    fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
+      unfolding less_Sup_iff by auto
+  qed
+next
+  assume *: "\<forall>x<top. \<exists>i\<in>A. x < i"
+  show "Sup A = top"
+  proof (rule ccontr)
+    assume "Sup A \<noteq> top"
+    with top_greatest[of "Sup A"]
+    have "Sup A < top" unfolding le_less by auto
+    then have "Sup A < Sup A"
+      using * unfolding less_Sup_iff by auto
+    then show False by auto
+  qed
+qed
+
+lemma SUP_eq_top_iff:
+  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
+  shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
+  unfolding SUPR_def Sup_eq_top_iff by auto
+
+lemma SUP_nat_Infty: "(SUP i::nat. ereal (real i)) = \<infinity>"
+proof -
+  { fix x ::ereal assume "x \<noteq> \<infinity>"
+    then have "\<exists>k::nat. x < ereal (real k)"
+    proof (cases x)
+      case MInf then show ?thesis by (intro exI[of _ 0]) auto
+    next
+      case (real r)
+      moreover obtain k :: nat where "r < real k"
+        using ex_less_of_nat by (auto simp: real_eq_of_nat)
+      ultimately show ?thesis by auto
+    qed simp }
+  then show ?thesis
+    using SUP_eq_top_iff[of UNIV "\<lambda>n::nat. ereal (real n)"]
+    by (auto simp: top_ereal_def)
+qed
+
+lemma ereal_le_Sup:
+  fixes x :: ereal
+  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
+(is "?lhs <-> ?rhs")
+proof-
+{ assume "?rhs"
+  { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
+    from this obtain y where y_def: "(SUP i:A. f i)<y & y<x" using ereal_dense by auto
+    from this obtain i where "i : A & y <= f i" using `?rhs` by auto
+    hence "y <= (SUP i:A. f i)" using le_SUPI[of i A f] by auto
+    hence False using y_def by auto
+  } hence "?lhs" by auto
+}
+moreover
+{ assume "?lhs" hence "?rhs"
+  by (metis Collect_def Collect_mem_eq SUP_leI assms atLeastatMost_empty atLeastatMost_empty_iff
+      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
+} ultimately show ?thesis by auto
+qed
+
+lemma ereal_Inf_le:
+  fixes x :: ereal
+  shows "((INF i:A. f i) <= x) <-> (ALL y. x < y --> (EX i. i : A & f i <= y))"
+(is "?lhs <-> ?rhs")
+proof-
+{ assume "?rhs"
+  { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le)
+    from this obtain y where y_def: "x<y & y<(INF i:A. f i)" using ereal_dense by auto
+    from this obtain i where "i : A & f i <= y" using `?rhs` by auto
+    hence "(INF i:A. f i) <= y" using INF_leI[of i A f] by auto
+    hence False using y_def by auto
+  } hence "?lhs" by auto
+}
+moreover
+{ assume "?lhs" hence "?rhs"
+  by (metis Collect_def Collect_mem_eq le_INFI assms atLeastatMost_empty atLeastatMost_empty_iff
+      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
+} ultimately show ?thesis by auto
+qed
+
+lemma Inf_less:
+  fixes x :: ereal
+  assumes "(INF i:A. f i) < x"
+  shows "EX i. i : A & f i <= x"
+proof(rule ccontr)
+  assume "~ (EX i. i : A & f i <= x)"
+  hence "ALL i:A. f i > x" by auto
+  hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto
+  thus False using assms by auto
+qed
+
+lemma same_INF:
+  assumes "ALL e:A. f e = g e"
+  shows "(INF e:A. f e) = (INF e:A. g e)"
+proof-
+have "f ` A = g ` A" unfolding image_def using assms by auto
+thus ?thesis unfolding INFI_def by auto
+qed
+
+lemma same_SUP:
+  assumes "ALL e:A. f e = g e"
+  shows "(SUP e:A. f e) = (SUP e:A. g e)"
+proof-
+have "f ` A = g ` A" unfolding image_def using assms by auto
+thus ?thesis unfolding SUPR_def by auto
+qed
+
+lemma SUPR_eq:
+  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<le> g j"
+  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<le> f i"
+  shows "(SUP i:A. f i) = (SUP j:B. g j)"
+proof (intro antisym)
+  show "(SUP i:A. f i) \<le> (SUP j:B. g j)"
+    using assms by (metis SUP_leI le_SUPI2)
+  show "(SUP i:B. g i) \<le> (SUP j:A. f j)"
+    using assms by (metis SUP_leI le_SUPI2)
+qed
+
+lemma SUP_ereal_le_addI:
+  fixes f :: "'i \<Rightarrow> ereal"
+  assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
+  shows "SUPR UNIV f + y \<le> z"
+proof (cases y)
+  case (real r)
+  then have "\<And>i. f i \<le> z - y" using assms by (simp add: ereal_le_minus_iff)
+  then have "SUPR UNIV f \<le> z - y" by (rule SUP_leI)
+  then show ?thesis using real by (simp add: ereal_le_minus_iff)
+qed (insert assms, auto)
+
+lemma SUPR_ereal_add:
+  fixes f g :: "nat \<Rightarrow> ereal"
+  assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
+  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
+proof (rule ereal_SUPI)
+  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
+  have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
+    unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD)
+  { fix j
+    { fix i
+      have "f i + g j \<le> f i + g (max i j)"
+        using `incseq g`[THEN incseqD] by (rule add_left_mono) auto
+      also have "\<dots> \<le> f (max i j) + g (max i j)"
+        using `incseq f`[THEN incseqD] by (rule add_right_mono) auto
+      also have "\<dots> \<le> y" using * by auto
+      finally have "f i + g j \<le> y" . }
+    then have "SUPR UNIV f + g j \<le> y"
+      using assms(4)[of j] by (intro SUP_ereal_le_addI) auto
+    then have "g j + SUPR UNIV f \<le> y" by (simp add: ac_simps) }
+  then have "SUPR UNIV g + SUPR UNIV f \<le> y"
+    using f by (rule SUP_ereal_le_addI)
+  then show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
+qed (auto intro!: add_mono le_SUPI)
+
+lemma SUPR_ereal_add_pos:
+  fixes f g :: "nat \<Rightarrow> ereal"
+  assumes inc: "incseq f" "incseq g" and pos: "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
+  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
+proof (intro SUPR_ereal_add inc)
+  fix i show "f i \<noteq> -\<infinity>" "g i \<noteq> -\<infinity>" using pos[of i] by auto
+qed
+
+lemma SUPR_ereal_setsum:
+  fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> ereal"
+  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)" and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
+  shows "(SUP i. \<Sum>n\<in>A. f n i) = (\<Sum>n\<in>A. SUPR UNIV (f n))"
+proof cases
+  assume "finite A" then show ?thesis using assms
+    by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_ereal_add_pos)
+qed simp
+
+lemma SUPR_ereal_cmult:
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
+  shows "(SUP i. c * f i) = c * SUPR UNIV f"
+proof (rule ereal_SUPI)
+  fix i have "f i \<le> SUPR UNIV f" by (rule le_SUPI) auto
+  then show "c * f i \<le> c * SUPR UNIV f"
+    using `0 \<le> c` by (rule ereal_mult_left_mono)
+next
+  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c * f i \<le> y"
+  show "c * SUPR UNIV f \<le> y"
+  proof cases
+    assume c: "0 < c \<and> c \<noteq> \<infinity>"
+    with * have "SUPR UNIV f \<le> y / c"
+      by (intro SUP_leI) (auto simp: ereal_le_divide_pos)
+    with c show ?thesis
+      by (auto simp: ereal_le_divide_pos)
+  next
+    { assume "c = \<infinity>" have ?thesis
+      proof cases
+        assume "\<forall>i. f i = 0"
+        moreover then have "range f = {0}" by auto
+        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
+      next
+        assume "\<not> (\<forall>i. f i = 0)"
+        then obtain i where "f i \<noteq> 0" by auto
+        with *[of i] `c = \<infinity>` `0 \<le> f i` show ?thesis by (auto split: split_if_asm)
+      qed }
+    moreover assume "\<not> (0 < c \<and> c \<noteq> \<infinity>)"
+    ultimately show ?thesis using * `0 \<le> c` by auto
+  qed
+qed
+
+lemma SUP_PInfty:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes "\<And>n::nat. \<exists>i\<in>A. ereal (real n) \<le> f i"
+  shows "(SUP i:A. f i) = \<infinity>"
+  unfolding SUPR_def Sup_eq_top_iff[where 'a=ereal, unfolded top_ereal_def]
+  apply simp
+proof safe
+  fix x :: ereal assume "x \<noteq> \<infinity>"
+  show "\<exists>i\<in>A. x < f i"
+  proof (cases x)
+    case PInf with `x \<noteq> \<infinity>` show ?thesis by simp
+  next
+    case MInf with assms[of "0"] show ?thesis by force
+  next
+    case (real r)
+    with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < ereal (real n)" by auto
+    moreover from assms[of n] guess i ..
+    ultimately show ?thesis
+      by (auto intro!: bexI[of _ i])
+  qed
+qed
+
+lemma Sup_countable_SUPR:
+  assumes "A \<noteq> {}"
+  shows "\<exists>f::nat \<Rightarrow> ereal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
+proof (cases "Sup A")
+  case (real r)
+  have "\<forall>n::nat. \<exists>x. x \<in> A \<and> Sup A < x + 1 / ereal (real n)"
+  proof
+    fix n ::nat have "\<exists>x\<in>A. Sup A - 1 / ereal (real n) < x"
+      using assms real by (intro Sup_ereal_close) (auto simp: one_ereal_def)
+    then guess x ..
+    then show "\<exists>x. x \<in> A \<and> Sup A < x + 1 / ereal (real n)"
+      by (auto intro!: exI[of _ x] simp: ereal_minus_less_iff)
+  qed
+  from choice[OF this] guess f .. note f = this
+  have "SUPR UNIV f = Sup A"
+  proof (rule ereal_SUPI)
+    fix i show "f i \<le> Sup A" using f
+      by (auto intro!: complete_lattice_class.Sup_upper)
+  next
+    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
+    show "Sup A \<le> y"
+    proof (rule ereal_le_epsilon, intro allI impI)
+      fix e :: ereal assume "0 < e"
+      show "Sup A \<le> y + e"
+      proof (cases e)
+        case (real r)
+        hence "0 < r" using `0 < e` by auto
+        then obtain n ::nat where *: "1 / real n < r" "0 < n"
+          using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
+        have "Sup A \<le> f n + 1 / ereal (real n)" using f[THEN spec, of n] by auto
+        also have "1 / ereal (real n) \<le> e" using real * by (auto simp: one_ereal_def )
+        with bound have "f n + 1 / ereal (real n) \<le> y + e" by (rule add_mono) simp
+        finally show "Sup A \<le> y + e" .
+      qed (insert `0 < e`, auto)
+    qed
+  qed
+  with f show ?thesis by (auto intro!: exI[of _ f])
+next
+  case PInf
+  from `A \<noteq> {}` obtain x where "x \<in> A" by auto
+  show ?thesis
+  proof cases
+    assume "\<infinity> \<in> A"
+    moreover then have "\<infinity> \<le> Sup A" by (intro complete_lattice_class.Sup_upper)
+    ultimately show ?thesis by (auto intro!: exI[of _ "\<lambda>x. \<infinity>"])
+  next
+    assume "\<infinity> \<notin> A"
+    have "\<exists>x\<in>A. 0 \<le> x"
+      by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least ereal_infty_less_eq2 linorder_linear)
+    then obtain x where "x \<in> A" "0 \<le> x" by auto
+    have "\<forall>n::nat. \<exists>f. f \<in> A \<and> x + ereal (real n) \<le> f"
+    proof (rule ccontr)
+      assume "\<not> ?thesis"
+      then have "\<exists>n::nat. Sup A \<le> x + ereal (real n)"
+        by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le)
+      then show False using `x \<in> A` `\<infinity> \<notin> A` PInf
+        by(cases x) auto
+    qed
+    from choice[OF this] guess f .. note f = this
+    have "SUPR UNIV f = \<infinity>"
+    proof (rule SUP_PInfty)
+      fix n :: nat show "\<exists>i\<in>UNIV. ereal (real n) \<le> f i"
+        using f[THEN spec, of n] `0 \<le> x`
+        by (cases rule: ereal2_cases[of "f n" x]) (auto intro!: exI[of _ n])
+    qed
+    then show ?thesis using f PInf by (auto intro!: exI[of _ f])
+  qed
+next
+  case MInf
+  with `A \<noteq> {}` have "A = {-\<infinity>}" by (auto simp: Sup_eq_MInfty)
+  then show ?thesis using MInf by (auto intro!: exI[of _ "\<lambda>x. -\<infinity>"])
+qed
+
+lemma SUPR_countable_SUPR:
+  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ereal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
+  using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def)
+
+lemma Sup_ereal_cadd:
+  fixes A :: "ereal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "Sup ((\<lambda>x. a + x) ` A) = a + Sup A"
+proof (rule antisym)
+  have *: "\<And>a::ereal. \<And>A. Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A"
+    by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
+  then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
+  show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
+  proof (cases a)
+    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
+  next
+    case (real r)
+    then have **: "op + (- a) ` op + a ` A = A"
+      by (auto simp: image_iff ac_simps zero_ereal_def[symmetric])
+    from *[of "-a" "(\<lambda>x. a + x) ` A"] real show ?thesis unfolding **
+      by (cases rule: ereal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto
+  qed (insert `a \<noteq> -\<infinity>`, auto)
+qed
+
+lemma Sup_ereal_cminus:
+  fixes A :: "ereal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "Sup ((\<lambda>x. a - x) ` A) = a - Inf A"
+  using Sup_ereal_cadd[of "uminus ` A" a] assms
+  by (simp add: comp_def image_image minus_ereal_def
+                 ereal_Sup_uminus_image_eq)
+
+lemma SUPR_ereal_cminus:
+  fixes f :: "'i \<Rightarrow> ereal"
+  fixes A assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
+  shows "(SUP x:A. a - f x) = a - (INF x:A. f x)"
+  using Sup_ereal_cminus[of "f`A" a] assms
+  unfolding SUPR_def INFI_def image_image by auto
+
+lemma Inf_ereal_cminus:
+  fixes A :: "ereal set" assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
+  shows "Inf ((\<lambda>x. a - x) ` A) = a - Sup A"
+proof -
+  { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto }
+  moreover then have "(\<lambda>x. -a - x)`uminus`A = uminus ` (\<lambda>x. a - x) ` A"
+    by (auto simp: image_image)
+  ultimately show ?thesis
+    using Sup_ereal_cminus[of "uminus ` A" "-a"] assms
+    by (auto simp add: ereal_Sup_uminus_image_eq ereal_Inf_uminus_image_eq)
+qed
+
+lemma INFI_ereal_cminus:
+  fixes a :: ereal assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
+  shows "(INF x:A. a - f x) = a - (SUP x:A. f x)"
+  using Inf_ereal_cminus[of "f`A" a] assms
+  unfolding SUPR_def INFI_def image_image
+  by auto
+
+lemma uminus_ereal_add_uminus_uminus:
+  fixes a b :: ereal shows "a \<noteq> \<infinity> \<Longrightarrow> b \<noteq> \<infinity> \<Longrightarrow> - (- a + - b) = a + b"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma INFI_ereal_add:
+  fixes f :: "nat \<Rightarrow> ereal"
+  assumes "decseq f" "decseq g" and fin: "\<And>i. f i \<noteq> \<infinity>" "\<And>i. g i \<noteq> \<infinity>"
+  shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g"
+proof -
+  have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
+    using assms unfolding INF_less_iff by auto
+  { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i"
+      by (rule uminus_ereal_add_uminus_uminus) }
+  then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
+    by simp
+  also have "\<dots> = INFI UNIV f + INFI UNIV g"
+    unfolding ereal_INFI_uminus
+    using assms INF_less
+    by (subst SUPR_ereal_add)
+       (auto simp: ereal_SUPR_uminus intro!: uminus_ereal_add_uminus_uminus)
+  finally show ?thesis .
+qed
+
+subsection "Limits on @{typ ereal}"
+
+subsubsection "Topological space"
+
+instantiation ereal :: topological_space
+begin
+
+definition "open A \<longleftrightarrow> open (ereal -` A)
+       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {ereal x <..} \<subseteq> A))
+       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A))"
+
+lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {ereal x<..} \<subseteq> A)"
+  unfolding open_ereal_def by auto
+
+lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A)"
+  unfolding open_ereal_def by auto
+
+lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{ereal x<..} \<subseteq> A"
+  using open_PInfty[OF assms] by auto
+
+lemma open_MInfty2: assumes "open A" "-\<infinity> \<in> A" obtains x where "{..<ereal x} \<subseteq> A"
+  using open_MInfty[OF assms] by auto
+
+lemma ereal_openE: assumes "open A" obtains x y where
+  "open (ereal -` A)"
+  "\<infinity> \<in> A \<Longrightarrow> {ereal x<..} \<subseteq> A"
+  "-\<infinity> \<in> A \<Longrightarrow> {..<ereal y} \<subseteq> A"
+  using assms open_ereal_def by auto
+
+instance
+proof
+  let ?U = "UNIV::ereal set"
+  show "open ?U" unfolding open_ereal_def
+    by (auto intro!: exI[of _ 0])
+next
+  fix S T::"ereal set" assume "open S" and "open T"
+  from `open S`[THEN ereal_openE] guess xS yS .
+  moreover from `open T`[THEN ereal_openE] guess xT yT .
+  ultimately have
+    "open (ereal -` (S \<inter> T))"
+    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {ereal (max xS xT) <..} \<subseteq> S \<inter> T"
+    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< ereal (min yS yT)} \<subseteq> S \<inter> T"
+    by auto
+  then show "open (S Int T)" unfolding open_ereal_def by blast
+next
+  fix K :: "ereal set set" assume "\<forall>S\<in>K. open S"
+  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (ereal -` S) \<and>
+    (\<infinity> \<in> S \<longrightarrow> {ereal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< ereal y} \<subseteq> S)"
+    by (auto simp: open_ereal_def)
+  then show "open (Union K)" unfolding open_ereal_def
+  proof (intro conjI impI)
+    show "open (ereal -` \<Union>K)"
+      using *[THEN choice] by (auto simp: vimage_Union)
+  qed ((metis UnionE Union_upper subset_trans *)+)
+qed
+end
+
+lemma open_ereal: "open S \<Longrightarrow> open (ereal ` S)"
+  by (auto simp: inj_vimage_image_eq open_ereal_def)
+
+lemma open_ereal_vimage: "open S \<Longrightarrow> open (ereal -` S)"
+  unfolding open_ereal_def by auto
+
+lemma open_ereal_lessThan[intro, simp]: "open {..< a :: ereal}"
+proof -
+  have "\<And>x. ereal -` {..<ereal x} = {..< x}"
+    "ereal -` {..< \<infinity>} = UNIV" "ereal -` {..< -\<infinity>} = {}" by auto
+  then show ?thesis by (cases a) (auto simp: open_ereal_def)
+qed
+
+lemma open_ereal_greaterThan[intro, simp]:
+  "open {a :: ereal <..}"
+proof -
+  have "\<And>x. ereal -` {ereal x<..} = {x<..}"
+    "ereal -` {\<infinity><..} = {}" "ereal -` {-\<infinity><..} = UNIV" by auto
+  then show ?thesis by (cases a) (auto simp: open_ereal_def)
+qed
+
+lemma ereal_open_greaterThanLessThan[intro, simp]: "open {a::ereal <..< b}"
+  unfolding greaterThanLessThan_def by auto
+
+lemma closed_ereal_atLeast[simp, intro]: "closed {a :: ereal ..}"
+proof -
+  have "- {a ..} = {..< a}" by auto
+  then show "closed {a ..}"
+    unfolding closed_def using open_ereal_lessThan by auto
+qed
+
+lemma closed_ereal_atMost[simp, intro]: "closed {.. b :: ereal}"
+proof -
+  have "- {.. b} = {b <..}" by auto
+  then show "closed {.. b}"
+    unfolding closed_def using open_ereal_greaterThan by auto
+qed
+
+lemma closed_ereal_atLeastAtMost[simp, intro]:
+  shows "closed {a :: ereal .. b}"
+  unfolding atLeastAtMost_def by auto
+
+lemma closed_ereal_singleton:
+  "closed {a :: ereal}"
+by (metis atLeastAtMost_singleton closed_ereal_atLeastAtMost)
+
+lemma ereal_open_cont_interval:
+  fixes S :: "ereal set"
+  assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
+  obtains e where "e>0" "{x-e <..< x+e} \<subseteq> S"
+proof-
+  from `open S` have "open (ereal -` S)" by (rule ereal_openE)
+  then obtain e where "0 < e" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> ereal y \<in> S"
+    using assms unfolding open_dist by force
+  show thesis
+  proof (intro that subsetI)
+    show "0 < ereal e" using `0 < e` by auto
+    fix y assume "y \<in> {x - ereal e<..<x + ereal e}"
+    with assms obtain t where "y = ereal t" "dist t (real x) < e"
+      apply (cases y) by (auto simp: dist_real_def)
+    then show "y \<in> S" using e[of t] by auto
+  qed
+qed
+
+lemma ereal_open_cont_interval2:
+  fixes S :: "ereal set"
+  assumes "open S" "x \<in> S" and x: "\<bar>x\<bar> \<noteq> \<infinity>"
+  obtains a b where "a < x" "x < b" "{a <..< b} \<subseteq> S"
+proof-
+  guess e using ereal_open_cont_interval[OF assms] .
+  with that[of "x-e" "x+e"] ereal_between[OF x, of e]
+  show thesis by auto
+qed
+
+instance ereal :: t2_space
+proof
+  fix x y :: ereal assume "x ~= y"
+  let "?P x (y::ereal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
+
+  { fix x y :: ereal assume "x < y"
+    from ereal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
+    have "?P x y"
+      apply (rule exI[of _ "{..<z}"])
+      apply (rule exI[of _ "{z<..}"])
+      using z by auto }
+  note * = this
+
+  from `x ~= y`
+  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
+  proof (cases rule: linorder_cases)
+    assume "x = y" with `x ~= y` show ?thesis by simp
+  next assume "x < y" from *[OF this] show ?thesis by auto
+  next assume "y < x" from *[OF this] show ?thesis by auto
+  qed
+qed
+
+subsubsection {* Convergent sequences *}
+
+lemma lim_ereal[simp]:
+  "((\<lambda>n. ereal (f n)) ---> ereal x) net \<longleftrightarrow> (f ---> x) net" (is "?l = ?r")
+proof (intro iffI topological_tendstoI)
+  fix S assume "?l" "open S" "x \<in> S"
+  then show "eventually (\<lambda>x. f x \<in> S) net"
+    using `?l`[THEN topological_tendstoD, OF open_ereal, OF `open S`]
+    by (simp add: inj_image_mem_iff)
+next
+  fix S assume "?r" "open S" "ereal x \<in> S"
+  show "eventually (\<lambda>x. ereal (f x) \<in> S) net"
+    using `?r`[THEN topological_tendstoD, OF open_ereal_vimage, OF `open S`]
+    using `ereal x \<in> S` by auto
+qed
+
+lemma lim_real_of_ereal[simp]:
+  assumes lim: "(f ---> ereal x) net"
+  shows "((\<lambda>x. real (f x)) ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume "open S" "x \<in> S"
+  then have S: "open S" "ereal x \<in> ereal ` S"
+    by (simp_all add: inj_image_mem_iff)
+  have "\<forall>x. f x \<in> ereal ` S \<longrightarrow> real (f x) \<in> S" by auto
+  from this lim[THEN topological_tendstoD, OF open_ereal, OF S]
+  show "eventually (\<lambda>x. real (f x) \<in> S) net"
+    by (rule eventually_mono)
+qed
+
+lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= ereal B)" (is "?l = ?r")
+proof
+  assume ?r
+  show ?l
+    apply(rule topological_tendstoI)
+    unfolding eventually_sequentially
+  proof-
+    fix S :: "ereal set" assume "open S" "\<infinity> : S"
+    from open_PInfty[OF this] guess B .. note B=this
+    from `?r`[rule_format,of "B+1"] guess N .. note N=this
+    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
+    proof safe case goal1
+      have "ereal B < ereal (B + 1)" by auto
+      also have "... <= f n" using goal1 N by auto
+      finally show ?case using B by fastsimp
+    qed
+  qed
+next
+  assume ?l
+  show ?r
+  proof fix B::real have "open {ereal B<..}" "\<infinity> : {ereal B<..}" by auto
+    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
+    guess N .. note N=this
+    show "EX N. ALL n>=N. ereal B <= f n" apply(rule_tac x=N in exI) using N by auto
+  qed
+qed
+
+
+lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= ereal B)" (is "?l = ?r")
+proof
+  assume ?r
+  show ?l
+    apply(rule topological_tendstoI)
+    unfolding eventually_sequentially
+  proof-
+    fix S :: "ereal set"
+    assume "open S" "(-\<infinity>) : S"
+    from open_MInfty[OF this] guess B .. note B=this
+    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
+    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
+    proof safe case goal1
+      have "ereal (B - 1) >= f n" using goal1 N by auto
+      also have "... < ereal B" by auto
+      finally show ?case using B by fastsimp
+    qed
+  qed
+next assume ?l show ?r
+  proof fix B::real have "open {..<ereal B}" "(-\<infinity>) : {..<ereal B}" by auto
+    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
+    guess N .. note N=this
+    show "EX N. ALL n>=N. ereal B >= f n" apply(rule_tac x=N in exI) using N by auto
+  qed
+qed
+
+
+lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= ereal B" shows "l ~= \<infinity>"
+proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
+  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
+  guess N .. note N=this[rule_format,OF le_refl]
+  hence "ereal ?B <= ereal B" using assms(2)[of N] by(rule order_trans)
+  hence "ereal ?B < ereal ?B" apply (rule le_less_trans) by auto
+  thus False by auto
+qed
+
+
+lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= ereal B" shows "l ~= (-\<infinity>)"
+proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
+  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
+  guess N .. note N=this[rule_format,OF le_refl]
+  hence "ereal B <= ereal ?B" using assms(2)[of N] order_trans[of "ereal B" "f N" "ereal(B - 1)"] by blast
+  thus False by auto
+qed
+
+
+lemma tendsto_explicit:
+  "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
+  unfolding tendsto_def eventually_sequentially by auto
+
+
+lemma tendsto_obtains_N:
+  assumes "f ----> f0"
+  assumes "open S" "f0 : S"
+  obtains N where "ALL n>=N. f n : S"
+  using tendsto_explicit[of f f0] assms by auto
+
+
+lemma tail_same_limit:
+  fixes X Y N
+  assumes "X ----> L" "ALL n>=N. X n = Y n"
+  shows "Y ----> L"
+proof-
+{ fix S assume "open S" and "L:S"
+  from this obtain N1 where "ALL n>=N1. X n : S"
+     using assms unfolding tendsto_def eventually_sequentially by auto
+  hence "ALL n>=max N N1. Y n : S" using assms by auto
+  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
+}
+thus ?thesis using tendsto_explicit by auto
+qed
+
+
+lemma Lim_bounded_PInfty2:
+assumes lim:"f ----> l" and "ALL n>=N. f n <= ereal B"
+shows "l ~= \<infinity>"
+proof-
+  def g == "(%n. if n>=N then f n else ereal B)"
+  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
+  moreover have "!!n. g n <= ereal B" using g_def assms by auto
+  ultimately show ?thesis using  Lim_bounded_PInfty by auto
+qed
+
+lemma Lim_bounded_ereal:
+  assumes lim:"f ----> (l :: ereal)"
+  and "ALL n>=M. f n <= C"
+  shows "l<=C"
+proof-
+{ assume "l=(-\<infinity>)" hence ?thesis by auto }
+moreover
+{ assume "~(l=(-\<infinity>))"
+  { assume "C=\<infinity>" hence ?thesis by auto }
+  moreover
+  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
+    hence "l=(-\<infinity>)" using assms
+       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
+    hence ?thesis by auto }
+  moreover
+  { assume "EX B. C = ereal B"
+    from this obtain B where B_def: "C=ereal B" by auto
+    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
+    from this obtain m where m_def: "ereal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
+    from this obtain N where N_def: "ALL n>=N. f n : {ereal(m - 1) <..< ereal(m+1)}"
+       apply (subst tendsto_obtains_N[of f l "{ereal(m - 1) <..< ereal(m+1)}"]) using assms by auto
+    { fix n assume "n>=N"
+      hence "EX r. ereal r = f n" using N_def by (cases "f n") auto
+    } from this obtain g where g_def: "ALL n>=N. ereal (g n) = f n" by metis
+    hence "(%n. ereal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
+    hence *: "(%n. g n) ----> m" using m_def by auto
+    { fix n assume "n>=max N M"
+      hence "ereal (g n) <= ereal B" using assms g_def B_def by auto
+      hence "g n <= B" by auto
+    } hence "EX N. ALL n>=N. g n <= B" by blast
+    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
+    hence ?thesis using m_def B_def by auto
+  } ultimately have ?thesis by (cases C) auto
+} ultimately show ?thesis by blast
+qed
+
+lemma real_of_ereal_mult[simp]:
+  fixes a b :: ereal shows "real (a * b) = real a * real b"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma real_of_ereal_eq_0:
+  fixes x :: ereal shows "real x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity> \<or> x = 0"
+  by (cases x) auto
+
+lemma tendsto_ereal_realD:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes "x \<noteq> 0" and tendsto: "((\<lambda>x. ereal (real (f x))) ---> x) net"
+  shows "(f ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume S: "open S" "x \<in> S"
+  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}" by auto
+  from tendsto[THEN topological_tendstoD, OF this]
+  show "eventually (\<lambda>x. f x \<in> S) net"
+    by (rule eventually_rev_mp) (auto simp: ereal_real real_of_ereal_0)
+qed
+
+lemma tendsto_ereal_realI:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>" and tendsto: "(f ---> x) net"
+  shows "((\<lambda>x. ereal (real (f x))) ---> x) net"
+proof (intro topological_tendstoI)
+  fix S assume "open S" "x \<in> S"
+  with x have "open (S - {\<infinity>, -\<infinity>})" "x \<in> S - {\<infinity>, -\<infinity>}" by auto
+  from tendsto[THEN topological_tendstoD, OF this]
+  show "eventually (\<lambda>x. ereal (real (f x)) \<in> S) net"
+    by (elim eventually_elim1) (auto simp: ereal_real)
+qed
+
+lemma ereal_mult_cancel_left:
+  fixes a b c :: ereal shows "a * b = a * c \<longleftrightarrow>
+    ((\<bar>a\<bar> = \<infinity> \<and> 0 < b * c) \<or> a = 0 \<or> b = c)"
+  by (cases rule: ereal3_cases[of a b c])
+     (simp_all add: zero_less_mult_iff)
+
+lemma ereal_inj_affinity:
+  fixes m t :: ereal
+  assumes "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" "\<bar>t\<bar> \<noteq> \<infinity>"
+  shows "inj_on (\<lambda>x. m * x + t) A"
+  using assms
+  by (cases rule: ereal2_cases[of m t])
+     (auto intro!: inj_onI simp: ereal_add_cancel_right ereal_mult_cancel_left)
+
+lemma ereal_PInfty_eq_plus[simp]:
+  fixes a b :: ereal
+  shows "\<infinity> = a + b \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_MInfty_eq_plus[simp]:
+  fixes a b :: ereal
+  shows "-\<infinity> = a + b \<longleftrightarrow> (a = -\<infinity> \<and> b \<noteq> \<infinity>) \<or> (b = -\<infinity> \<and> a \<noteq> \<infinity>)"
+  by (cases rule: ereal2_cases[of a b]) auto
+
+lemma ereal_less_divide_pos:
+  fixes x y :: ereal
+  shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y < z / x \<longleftrightarrow> x * y < z"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_divide_less_pos:
+  fixes x y z :: ereal
+  shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y / x < z \<longleftrightarrow> y < x * z"
+  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
+
+lemma ereal_divide_eq:
+  fixes a b c :: ereal
+  shows "b \<noteq> 0 \<Longrightarrow> \<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> a / b = c \<longleftrightarrow> a = b * c"
+  by (cases rule: ereal3_cases[of a b c])
+     (simp_all add: field_simps)
+
+lemma ereal_inverse_not_MInfty[simp]: "inverse (a::ereal) \<noteq> -\<infinity>"
+  by (cases a) auto
+
+lemma ereal_mult_m1[simp]: "x * ereal (-1) = -x"
+  by (cases x) auto
+
+lemma ereal_LimI_finite:
+  fixes x :: ereal
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
+  assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r"
+  shows "u ----> x"
+proof (rule topological_tendstoI, unfold eventually_sequentially)
+  obtain rx where rx_def: "x=ereal rx" using assms by (cases x) auto
+  fix S assume "open S" "x : S"
+  then have "open (ereal -` S)" unfolding open_ereal_def by auto
+  with `x \<in> S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> ereal y \<in> S"
+    unfolding open_real_def rx_def by auto
+  then obtain n where
+    upper: "!!N. n <= N ==> u N < x + ereal r" and
+    lower: "!!N. n <= N ==> x < u N + ereal r" using assms(2)[of "ereal r"] by auto
+  show "EX N. ALL n>=N. u n : S"
+  proof (safe intro!: exI[of _ n])
+    fix N assume "n <= N"
+    from upper[OF this] lower[OF this] assms `0 < r`
+    have "u N ~: {\<infinity>,(-\<infinity>)}" by auto
+    from this obtain ra where ra_def: "(u N) = ereal ra" by (cases "u N") auto
+    hence "rx < ra + r" and "ra < rx + r"
+       using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto
+    hence "dist (real (u N)) rx < r"
+      using rx_def ra_def
+      by (auto simp: dist_real_def abs_diff_less_iff field_simps)
+    from dist[OF this] show "u N : S" using `u N  ~: {\<infinity>, -\<infinity>}`
+      by (auto simp: ereal_real split: split_if_asm)
+  qed
+qed
+
+lemma ereal_LimI_finite_iff:
+  fixes x :: ereal
+  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
+  shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))"
+  (is "?lhs <-> ?rhs")
+proof
+  assume lim: "u ----> x"
+  { fix r assume "(r::ereal)>0"
+    from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}"
+       apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
+       using lim ereal_between[of x r] assms `r>0` by auto
+    hence "EX N. ALL n>=N. u n < x + r & x < u n + r"
+      using ereal_minus_less[of r x] by (cases r) auto
+  } then show "?rhs" by auto
+next
+  assume ?rhs then show "u ----> x"
+    using ereal_LimI_finite[of x] assms by auto
+qed
+
+
+subsubsection {* @{text Liminf} and @{text Limsup} *}
+
+definition
+  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
+
+definition
+  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
+
+lemma Liminf_Sup:
+  fixes f :: "'a => 'b::{complete_lattice, linorder}"
+  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
+  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
+
+lemma Limsup_Inf:
+  fixes f :: "'a => 'b::{complete_lattice, linorder}"
+  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
+  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
+
+lemma ereal_SupI:
+  fixes x :: ereal
+  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
+  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
+  shows "Sup A = x"
+  unfolding Sup_ereal_def
+  using assms by (auto intro!: Least_equality)
+
+lemma ereal_InfI:
+  fixes x :: ereal
+  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
+  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
+  shows "Inf A = x"
+  unfolding Inf_ereal_def
+  using assms by (auto intro!: Greatest_equality)
+
+lemma Limsup_const:
+  fixes c :: "'a::{complete_lattice, linorder}"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Limsup net (\<lambda>x. c) = c"
+  unfolding Limsup_Inf
+proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
+  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
+  show "c \<le> x"
+  proof (rule ccontr)
+    assume "\<not> c \<le> x" then have "x < c" by auto
+    then show False using ntriv * by (auto simp: trivial_limit_def)
+  qed
+qed auto
+
+lemma Liminf_const:
+  fixes c :: "'a::{complete_lattice, linorder}"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Liminf net (\<lambda>x. c) = c"
+  unfolding Liminf_Sup
+proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
+  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
+  show "x \<le> c"
+  proof (rule ccontr)
+    assume "\<not> x \<le> c" then have "c < x" by auto
+    then show False using ntriv * by (auto simp: trivial_limit_def)
+  qed
+qed auto
+
+lemma mono_set:
+  fixes S :: "('a::order) set"
+  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
+  by (auto simp: mono_def mem_def)
+
+lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
+lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
+lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
+lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
+
+lemma mono_set_iff:
+  fixes S :: "'a::{linorder,complete_lattice} set"
+  defines "a \<equiv> Inf S"
+  shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
+proof
+  assume "mono S"
+  then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set)
+  show ?c
+  proof cases
+    assume "a \<in> S"
+    show ?c
+      using mono[OF _ `a \<in> S`]
+      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
+  next
+    assume "a \<notin> S"
+    have "S = {a <..}"
+    proof safe
+      fix x assume "x \<in> S"
+      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
+      then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
+    next
+      fix x assume "a < x"
+      then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff ..
+      with mono[of y x] show "x \<in> S" by auto
+    qed
+    then show ?c ..
+  qed
+qed auto
+
+lemma lim_imp_Liminf:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes ntriv: "\<not> trivial_limit net"
+  assumes lim: "(f ---> f0) net"
+  shows "Liminf net f = f0"
+  unfolding Liminf_Sup
+proof (safe intro!: ereal_SupI)
+  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
+  show "y \<le> f0"
+  proof (rule ereal_le_ereal)
+    fix B assume "B < y"
+    { assume "f0 < B"
+      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
+         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
+         by (auto intro: eventually_conj)
+      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
+      finally have False using ntriv[unfolded trivial_limit_def] by auto
+    } then show "B \<le> f0" by (metis linorder_le_less_linear)
+  qed
+next
+  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
+  show "f0 \<le> y"
+  proof (safe intro!: *[rule_format])
+    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
+      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
+  qed
+qed
+
+lemma ereal_Liminf_le_Limsup:
+  fixes f :: "'a \<Rightarrow> ereal"
+  assumes ntriv: "\<not> trivial_limit net"
+  shows "Liminf net f \<le> Limsup net f"
+  unfolding Limsup_Inf Liminf_Sup
+proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
+  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
+  show "u \<le> v"
+  proof (rule ccontr)
+    assume "\<not> u \<le> v"
+    then obtain t where "t < u" "v < t"
+      using ereal_dense[of v u] by (auto simp: not_le)
+    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
+      using * by (auto intro: eventually_conj)
+    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
+    finally show False using ntriv by (auto simp: trivial_limit_def)
+  qed
+qed
+
+lemma Liminf_mono:
+  fixes f g :: "'a => ereal"
+  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
+  shows "Liminf net f \<le> Liminf net g"
+  unfolding Liminf_Sup
+proof (safe intro!: Sup_mono bexI)
+  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
+  then have "eventually (\<lambda>x. y < f x) net" by auto
+  then show "eventually (\<lambda>x. y < g x) net"
+    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
+qed simp
+
+lemma Liminf_eq:
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes "eventually (\<lambda>x. f x = g x) net"
+  shows "Liminf net f = Liminf net g"
+  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
+
+lemma Liminf_mono_all:
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes "\<And>x. f x \<le> g x"
+  shows "Liminf net f \<le> Liminf net g"
+  using assms by (intro Liminf_mono always_eventually) auto
+
+lemma Limsup_mono:
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
+  shows "Limsup net f \<le> Limsup net g"
+  unfolding Limsup_Inf
+proof (safe intro!: Inf_mono bexI)
+  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
+  then have "eventually (\<lambda>x. g x < y) net" by auto
+  then show "eventually (\<lambda>x. f x < y) net"
+    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
+qed simp
+
+lemma Limsup_mono_all:
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes "\<And>x. f x \<le> g x"
+  shows "Limsup net f \<le> Limsup net g"
+  using assms by (intro Limsup_mono always_eventually) auto
+
+lemma Limsup_eq:
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes "eventually (\<lambda>x. f x = g x) net"
+  shows "Limsup net f = Limsup net g"
+  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
+
+abbreviation "liminf \<equiv> Liminf sequentially"
+
+abbreviation "limsup \<equiv> Limsup sequentially"
+
+lemma (in complete_lattice) less_INFD:
+  assumes "y < INFI A f"" i \<in> A" shows "y < f i"
+proof -
+  note `y < INFI A f`
+  also have "INFI A f \<le> f i" using `i \<in> A` by (rule INF_leI)
+  finally show "y < f i" .
+qed
+
+lemma liminf_SUPR_INFI:
+  fixes f :: "nat \<Rightarrow> ereal"
+  shows "liminf f = (SUP n. INF m:{n..}. f m)"
+  unfolding Liminf_Sup eventually_sequentially
+proof (safe intro!: antisym complete_lattice_class.Sup_least)
+  fix x assume *: "\<forall>y<x. \<exists>N. \<forall>n\<ge>N. y < f n" show "x \<le> (SUP n. INF m:{n..}. f m)"
+  proof (rule ereal_le_ereal)
+    fix y assume "y < x"
+    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
+    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
+    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto
+    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
+  qed
+next
+  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
+  proof (unfold SUPR_def, safe intro!: Sup_mono bexI)
+    fix y n assume "y < INFI {n..} f"
+    from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
+  qed (rule order_refl)
+qed
+
+lemma tail_same_limsup:
+  fixes X Y :: "nat => ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
+  shows "limsup X = limsup Y"
+  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma tail_same_liminf:
+  fixes X Y :: "nat => ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
+  shows "liminf X = liminf Y"
+  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma liminf_mono:
+  fixes X Y :: "nat \<Rightarrow> ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
+  shows "liminf X \<le> liminf Y"
+  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
+
+lemma limsup_mono:
+  fixes X Y :: "nat => ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
+  shows "limsup X \<le> limsup Y"
+  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
+
+declare trivial_limit_sequentially[simp]
+
+lemma
+  fixes X :: "nat \<Rightarrow> ereal"
+  shows ereal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
+    and ereal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
+  unfolding incseq_def decseq_def by auto
+
+lemma liminf_bounded:
+  fixes X Y :: "nat \<Rightarrow> ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
+  shows "C \<le> liminf X"
+  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
+
+lemma limsup_bounded:
+  fixes X Y :: "nat => ereal"
+  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
+  shows "limsup X \<le> C"
+  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
+
+lemma liminf_bounded_iff:
+  fixes x :: "nat \<Rightarrow> ereal"
+  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
+proof safe
+  fix B assume "B < C" "C \<le> liminf x"
+  then have "B < liminf x" by auto
+  then obtain N where "B < (INF m:{N..}. x m)"
+    unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto
+  from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
+next
+  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
+  { fix B assume "B<C"
+    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
+    hence "B \<le> (INF m:{N..}. x m)" by (intro le_INFI) auto
+    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp
+    finally have "B \<le> liminf x" .
+  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
+qed
+
+lemma liminf_subseq_mono:
+  fixes X :: "nat \<Rightarrow> ereal"
+  assumes "subseq r"
+  shows "liminf X \<le> liminf (X \<circ> r) "
+proof-
+  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
+  proof (safe intro!: INF_mono)
+    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
+      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
+  qed
+  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
+qed
+
+lemma ereal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "ereal (real x) = x"
+  using assms by auto
+
+lemma ereal_le_ereal_bounded:
+  fixes x y z :: ereal
+  assumes "z \<le> y"
+  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
+  shows "x \<le> y"
+proof (rule ereal_le_ereal)
+  fix B assume "B < x"
+  show "B \<le> y"
+  proof cases
+    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
+  next
+    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
+  qed
+qed
+
+lemma fixes x y :: ereal
+  shows Sup_atMost[simp]: "Sup {.. y} = y"
+    and Sup_lessThan[simp]: "Sup {..< y} = y"
+    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
+    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
+    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
+  by (auto simp: Sup_ereal_def intro!: Least_equality
+           intro: ereal_le_ereal ereal_le_ereal_bounded[of x])
+
+lemma Sup_greaterThanlessThan[simp]:
+  fixes x y :: ereal assumes "x < y" shows "Sup { x <..< y} = y"
+  unfolding Sup_ereal_def
+proof (intro Least_equality ereal_le_ereal_bounded[of _ _ y])
+  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
+  from ereal_dense[OF `x < y`] guess w .. note w = this
+  with z[THEN bspec, of w] show "x \<le> z" by auto
+qed auto
+
+lemma real_ereal_id: "real o ereal = id"
+proof-
+{ fix x have "(real o ereal) x = id x" by auto }
+from this show ?thesis using ext by blast
+qed
+
+lemma open_image_ereal: "open(UNIV-{ \<infinity> , (-\<infinity> :: ereal)})"
+by (metis range_ereal open_ereal open_UNIV)
+
+lemma ereal_le_distrib:
+  fixes a b c :: ereal shows "c * (a + b) \<le> c * a + c * b"
+  by (cases rule: ereal3_cases[of a b c])
+     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
+
+lemma ereal_pos_distrib:
+  fixes a b c :: ereal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
+  using assms by (cases rule: ereal3_cases[of a b c])
+                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
+
+lemma ereal_pos_le_distrib:
+fixes a b c :: ereal
+assumes "c>=0"
+shows "c * (a + b) <= c * a + c * b"
+  using assms by (cases rule: ereal3_cases[of a b c])
+                 (auto simp add: field_simps)
+
+lemma ereal_max_mono:
+  "[| (a::ereal) <= b; c <= d |] ==> max a c <= max b d"
+  by (metis sup_ereal_def sup_mono)
+
+
+lemma ereal_max_least:
+  "[| (a::ereal) <= x; c <= x |] ==> max a c <= x"
+  by (metis sup_ereal_def sup_least)
+
+end
--- a/src/HOL/Library/Extended_Reals.thy	Wed Jul 20 10:11:08 2011 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2535 +0,0 @@
-(*  Title:      HOL/Library/Extended_Reals.thy
-    Author:     Johannes Hölzl, TU München
-    Author:     Robert Himmelmann, TU München
-    Author:     Armin Heller, TU München
-    Author:     Bogdan Grechuk, University of Edinburgh
-*)
-
-header {* Extended real number line *}
-
-theory Extended_Reals
-  imports Complex_Main
-begin
-
-text {*
-
-For more lemmas about the extended real numbers go to
-  @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"}
-
-*}
-
-lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
-proof
-  assume "{x..} = UNIV"
-  show "x = bot"
-  proof (rule ccontr)
-    assume "x \<noteq> bot" then have "bot \<notin> {x..}" by (simp add: le_less)
-    then show False using `{x..} = UNIV` by simp
-  qed
-qed auto
-
-lemma SUPR_pair:
-  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
-  by (rule antisym) (auto intro!: SUP_leI le_SUPI2)
-
-lemma INFI_pair:
-  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
-  by (rule antisym) (auto intro!: le_INFI INF_leI2)
-
-subsection {* Definition and basic properties *}
-
-datatype extreal = extreal real | PInfty | MInfty
-
-notation (xsymbols)
-  PInfty  ("\<infinity>")
-
-notation (HTML output)
-  PInfty  ("\<infinity>")
-
-declare [[coercion "extreal :: real \<Rightarrow> extreal"]]
-
-instantiation extreal :: uminus
-begin
-  fun uminus_extreal where
-    "- (extreal r) = extreal (- r)"
-  | "- \<infinity> = MInfty"
-  | "- MInfty = \<infinity>"
-  instance ..
-end
-
-lemma inj_extreal[simp]: "inj_on extreal A"
-  unfolding inj_on_def by auto
-
-lemma MInfty_neq_PInfty[simp]:
-  "\<infinity> \<noteq> - \<infinity>" "- \<infinity> \<noteq> \<infinity>" by simp_all
-
-lemma MInfty_neq_extreal[simp]:
-  "extreal r \<noteq> - \<infinity>" "- \<infinity> \<noteq> extreal r" by simp_all
-
-lemma MInfinity_cases[simp]:
-  "(case - \<infinity> of extreal r \<Rightarrow> f r | \<infinity> \<Rightarrow> y | MInfinity \<Rightarrow> z) = z"
-  by simp
-
-lemma extreal_uminus_uminus[simp]:
-  fixes a :: extreal shows "- (- a) = a"
-  by (cases a) simp_all
-
-lemma MInfty_eq[simp, code_post]:
-  "MInfty = - \<infinity>" by simp
-
-declare uminus_extreal.simps(2)[code_inline, simp del]
-
-lemma extreal_cases[case_names real PInf MInf, cases type: extreal]:
-  assumes "\<And>r. x = extreal r \<Longrightarrow> P"
-  assumes "x = \<infinity> \<Longrightarrow> P"
-  assumes "x = -\<infinity> \<Longrightarrow> P"
-  shows P
-  using assms by (cases x) auto
-
-lemmas extreal2_cases = extreal_cases[case_product extreal_cases]
-lemmas extreal3_cases = extreal2_cases[case_product extreal_cases]
-
-lemma extreal_uminus_eq_iff[simp]:
-  fixes a b :: extreal shows "-a = -b \<longleftrightarrow> a = b"
-  by (cases rule: extreal2_cases[of a b]) simp_all
-
-function of_extreal :: "extreal \<Rightarrow> real" where
-"of_extreal (extreal r) = r" |
-"of_extreal \<infinity> = 0" |
-"of_extreal (-\<infinity>) = 0"
-  by (auto intro: extreal_cases)
-termination proof qed (rule wf_empty)
-
-defs (overloaded)
-  real_of_extreal_def [code_unfold]: "real \<equiv> of_extreal"
-
-lemma real_of_extreal[simp]:
-    "real (- x :: extreal) = - (real x)"
-    "real (extreal r) = r"
-    "real \<infinity> = 0"
-  by (cases x) (simp_all add: real_of_extreal_def)
-
-lemma range_extreal[simp]: "range extreal = UNIV - {\<infinity>, -\<infinity>}"
-proof safe
-  fix x assume "x \<notin> range extreal" "x \<noteq> \<infinity>"
-  then show "x = -\<infinity>" by (cases x) auto
-qed auto
-
-lemma extreal_range_uminus[simp]: "range uminus = (UNIV::extreal set)"
-proof safe
-  fix x :: extreal show "x \<in> range uminus" by (intro image_eqI[of _ _ "-x"]) auto
-qed auto
-
-instantiation extreal :: number
-begin
-definition [simp]: "number_of x = extreal (number_of x)"
-instance proof qed
-end
-
-instantiation extreal :: abs
-begin
-  function abs_extreal where
-    "\<bar>extreal r\<bar> = extreal \<bar>r\<bar>"
-  | "\<bar>-\<infinity>\<bar> = \<infinity>"
-  | "\<bar>\<infinity>\<bar> = \<infinity>"
-  by (auto intro: extreal_cases)
-  termination proof qed (rule wf_empty)
-  instance ..
-end
-
-lemma abs_eq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> = \<infinity> ; x = \<infinity> \<Longrightarrow> P ; x = -\<infinity> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
-  by (cases x) auto
-
-lemma abs_neq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> \<noteq> \<infinity> ; \<And>r. x = extreal r \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
-  by (cases x) auto
-
-lemma abs_extreal_uminus[simp]: "\<bar>- x\<bar> = \<bar>x::extreal\<bar>"
-  by (cases x) auto
-
-subsubsection "Addition"
-
-instantiation extreal :: comm_monoid_add
-begin
-
-definition "0 = extreal 0"
-
-function plus_extreal where
-"extreal r + extreal p = extreal (r + p)" |
-"\<infinity> + a = \<infinity>" |
-"a + \<infinity> = \<infinity>" |
-"extreal r + -\<infinity> = - \<infinity>" |
-"-\<infinity> + extreal p = -\<infinity>" |
-"-\<infinity> + -\<infinity> = -\<infinity>"
-proof -
-  case (goal1 P x)
-  moreover then obtain a b where "x = (a, b)" by (cases x) auto
-  ultimately show P
-   by (cases rule: extreal2_cases[of a b]) auto
-qed auto
-termination proof qed (rule wf_empty)
-
-lemma Infty_neq_0[simp]:
-  "\<infinity> \<noteq> 0" "0 \<noteq> \<infinity>"
-  "-\<infinity> \<noteq> 0" "0 \<noteq> -\<infinity>"
-  by (simp_all add: zero_extreal_def)
-
-lemma extreal_eq_0[simp]:
-  "extreal r = 0 \<longleftrightarrow> r = 0"
-  "0 = extreal r \<longleftrightarrow> r = 0"
-  unfolding zero_extreal_def by simp_all
-
-instance
-proof
-  fix a :: extreal show "0 + a = a"
-    by (cases a) (simp_all add: zero_extreal_def)
-  fix b :: extreal show "a + b = b + a"
-    by (cases rule: extreal2_cases[of a b]) simp_all
-  fix c :: extreal show "a + b + c = a + (b + c)"
-    by (cases rule: extreal3_cases[of a b c]) simp_all
-qed
-end
-
-lemma real_of_extreal_0[simp]: "real (0::extreal) = 0"
-  unfolding real_of_extreal_def zero_extreal_def by simp
-
-lemma abs_extreal_zero[simp]: "\<bar>0\<bar> = (0::extreal)"
-  unfolding zero_extreal_def abs_extreal.simps by simp
-
-lemma extreal_uminus_zero[simp]:
-  "- 0 = (0::extreal)"
-  by (simp add: zero_extreal_def)
-
-lemma extreal_uminus_zero_iff[simp]:
-  fixes a :: extreal shows "-a = 0 \<longleftrightarrow> a = 0"
-  by (cases a) simp_all
-
-lemma extreal_plus_eq_PInfty[simp]:
-  shows "a + b = \<infinity> \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_plus_eq_MInfty[simp]:
-  shows "a + b = -\<infinity> \<longleftrightarrow>
-    (a = -\<infinity> \<or> b = -\<infinity>) \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_add_cancel_left:
-  assumes "a \<noteq> -\<infinity>"
-  shows "a + b = a + c \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
-  using assms by (cases rule: extreal3_cases[of a b c]) auto
-
-lemma extreal_add_cancel_right:
-  assumes "a \<noteq> -\<infinity>"
-  shows "b + a = c + a \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
-  using assms by (cases rule: extreal3_cases[of a b c]) auto
-
-lemma extreal_real:
-  "extreal (real x) = (if \<bar>x\<bar> = \<infinity> then 0 else x)"
-  by (cases x) simp_all
-
-lemma real_of_extreal_add:
-  fixes a b :: extreal
-  shows "real (a + b) = (if (\<bar>a\<bar> = \<infinity>) \<and> (\<bar>b\<bar> = \<infinity>) \<or> (\<bar>a\<bar> \<noteq> \<infinity>) \<and> (\<bar>b\<bar> \<noteq> \<infinity>) then real a + real b else 0)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-subsubsection "Linear order on @{typ extreal}"
-
-instantiation extreal :: linorder
-begin
-
-function less_extreal where
-"extreal x < extreal y \<longleftrightarrow> x < y" |
-"        \<infinity> < a         \<longleftrightarrow> False" |
-"        a < -\<infinity>        \<longleftrightarrow> False" |
-"extreal x < \<infinity>         \<longleftrightarrow> True" |
-"       -\<infinity> < extreal r \<longleftrightarrow> True" |
-"       -\<infinity> < \<infinity>         \<longleftrightarrow> True"
-proof -
-  case (goal1 P x)
-  moreover then obtain a b where "x = (a,b)" by (cases x) auto
-  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
-qed simp_all
-termination by (relation "{}") simp
-
-definition "x \<le> (y::extreal) \<longleftrightarrow> x < y \<or> x = y"
-
-lemma extreal_infty_less[simp]:
-  "x < \<infinity> \<longleftrightarrow> (x \<noteq> \<infinity>)"
-  "-\<infinity> < x \<longleftrightarrow> (x \<noteq> -\<infinity>)"
-  by (cases x, simp_all) (cases x, simp_all)
-
-lemma extreal_infty_less_eq[simp]:
-  "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
-  "x \<le> -\<infinity> \<longleftrightarrow> x = -\<infinity>"
-  by (auto simp add: less_eq_extreal_def)
-
-lemma extreal_less[simp]:
-  "extreal r < 0 \<longleftrightarrow> (r < 0)"
-  "0 < extreal r \<longleftrightarrow> (0 < r)"
-  "0 < \<infinity>"
-  "-\<infinity> < 0"
-  by (simp_all add: zero_extreal_def)
-
-lemma extreal_less_eq[simp]:
-  "x \<le> \<infinity>"
-  "-\<infinity> \<le> x"
-  "extreal r \<le> extreal p \<longleftrightarrow> r \<le> p"
-  "extreal r \<le> 0 \<longleftrightarrow> r \<le> 0"
-  "0 \<le> extreal r \<longleftrightarrow> 0 \<le> r"
-  by (auto simp add: less_eq_extreal_def zero_extreal_def)
-
-lemma extreal_infty_less_eq2:
-  "a \<le> b \<Longrightarrow> a = \<infinity> \<Longrightarrow> b = \<infinity>"
-  "a \<le> b \<Longrightarrow> b = -\<infinity> \<Longrightarrow> a = -\<infinity>"
-  by simp_all
-
-instance
-proof
-  fix x :: extreal show "x \<le> x"
-    by (cases x) simp_all
-  fix y :: extreal show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
-    by (cases rule: extreal2_cases[of x y]) auto
-  show "x \<le> y \<or> y \<le> x "
-    by (cases rule: extreal2_cases[of x y]) auto
-  { assume "x \<le> y" "y \<le> x" then show "x = y"
-    by (cases rule: extreal2_cases[of x y]) auto }
-  { fix z assume "x \<le> y" "y \<le> z" then show "x \<le> z"
-    by (cases rule: extreal3_cases[of x y z]) auto }
-qed
-end
-
-instance extreal :: ordered_ab_semigroup_add
-proof
-  fix a b c :: extreal assume "a \<le> b" then show "c + a \<le> c + b"
-    by (cases rule: extreal3_cases[of a b c]) auto
-qed
-
-lemma real_of_extreal_positive_mono:
-  "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
-  by (cases rule: extreal2_cases[of x y]) auto
-
-lemma extreal_MInfty_lessI[intro, simp]:
-  "a \<noteq> -\<infinity> \<Longrightarrow> -\<infinity> < a"
-  by (cases a) auto
-
-lemma extreal_less_PInfty[intro, simp]:
-  "a \<noteq> \<infinity> \<Longrightarrow> a < \<infinity>"
-  by (cases a) auto
-
-lemma extreal_less_extreal_Ex:
-  fixes a b :: extreal
-  shows "x < extreal r \<longleftrightarrow> x = -\<infinity> \<or> (\<exists>p. p < r \<and> x = extreal p)"
-  by (cases x) auto
-
-lemma less_PInf_Ex_of_nat: "x \<noteq> \<infinity> \<longleftrightarrow> (\<exists>n::nat. x < extreal (real n))"
-proof (cases x)
-  case (real r) then show ?thesis
-    using reals_Archimedean2[of r] by simp
-qed simp_all
-
-lemma extreal_add_mono:
-  fixes a b c d :: extreal assumes "a \<le> b" "c \<le> d" shows "a + c \<le> b + d"
-  using assms
-  apply (cases a)
-  apply (cases rule: extreal3_cases[of b c d], auto)
-  apply (cases rule: extreal3_cases[of b c d], auto)
-  done
-
-lemma extreal_minus_le_minus[simp]:
-  fixes a b :: extreal shows "- a \<le> - b \<longleftrightarrow> b \<le> a"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_minus_less_minus[simp]:
-  fixes a b :: extreal shows "- a < - b \<longleftrightarrow> b < a"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_le_real_iff:
-  "x \<le> real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x \<le> y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x \<le> 0))"
-  by (cases y) auto
-
-lemma real_le_extreal_iff:
-  "real y \<le> x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y \<le> extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 \<le> x))"
-  by (cases y) auto
-
-lemma extreal_less_real_iff:
-  "x < real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x < y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x < 0))"
-  by (cases y) auto
-
-lemma real_less_extreal_iff:
-  "real y < x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y < extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 < x))"
-  by (cases y) auto
-
-lemma real_of_extreal_pos:
-  fixes x :: extreal shows "0 \<le> x \<Longrightarrow> 0 \<le> real x" by (cases x) auto
-
-lemmas real_of_extreal_ord_simps =
-  extreal_le_real_iff real_le_extreal_iff extreal_less_real_iff real_less_extreal_iff
-
-lemma abs_extreal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: extreal\<bar> = x"
-  by (cases x) auto
-
-lemma abs_extreal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: extreal\<bar> = -x"
-  by (cases x) auto
-
-lemma abs_extreal_pos[simp]: "0 \<le> \<bar>x :: extreal\<bar>"
-  by (cases x) auto
-
-lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
-  by (cases X) auto
-
-lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
-  by (cases X) auto
-
-lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
-  by (cases X) auto
-
-lemma extreal_0_le_uminus_iff[simp]:
-  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
-  by (cases rule: extreal2_cases[of a]) auto
-
-lemma extreal_uminus_le_0_iff[simp]:
-  fixes a :: extreal shows "-a \<le> 0 \<longleftrightarrow> 0 \<le> a"
-  by (cases rule: extreal2_cases[of a]) auto
-
-lemma extreal_dense:
-  fixes x y :: extreal assumes "x < y"
-  shows "EX z. x < z & z < y"
-proof -
-{ assume a: "x = (-\<infinity>)"
-  { assume "y = \<infinity>" hence ?thesis using a by (auto intro!: exI[of _ "0"]) }
-  moreover
-  { assume "y ~= \<infinity>"
-    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
-    hence ?thesis using `x < y` a by (auto intro!: exI[of _ "extreal (r - 1)"])
-  } ultimately have ?thesis by auto
-}
-moreover
-{ assume "x ~= (-\<infinity>)"
-  with `x < y` obtain p where p: "x = extreal p" by (cases x) auto
-  { assume "y = \<infinity>" hence ?thesis using `x < y` p
-       by (auto intro!: exI[of _ "extreal (p + 1)"]) }
-  moreover
-  { assume "y ~= \<infinity>"
-    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
-    with p `x < y` have "p < r" by auto
-    with dense obtain z where "p < z" "z < r" by auto
-    hence ?thesis using r p by (auto intro!: exI[of _ "extreal z"])
-  } ultimately have ?thesis by auto
-} ultimately show ?thesis by auto
-qed
-
-lemma extreal_dense2:
-  fixes x y :: extreal assumes "x < y"
-  shows "EX z. x < extreal z & extreal z < y"
-  by (metis extreal_dense[OF `x < y`] extreal_cases less_extreal.simps(2,3))
-
-lemma extreal_add_strict_mono:
-  fixes a b c d :: extreal
-  assumes "a = b" "0 \<le> a" "a \<noteq> \<infinity>" "c < d"
-  shows "a + c < b + d"
-  using assms by (cases rule: extreal3_cases[case_product extreal_cases, of a b c d]) auto
-
-lemma extreal_less_add: "\<bar>a\<bar> \<noteq> \<infinity> \<Longrightarrow> c < b \<Longrightarrow> a + c < a + b"
-  by (cases rule: extreal2_cases[of b c]) auto
-
-lemma extreal_uminus_eq_reorder: "- a = b \<longleftrightarrow> a = (-b::extreal)" by auto
-
-lemma extreal_uminus_less_reorder: "- a < b \<longleftrightarrow> -b < (a::extreal)"
-  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_less_minus)
-
-lemma extreal_uminus_le_reorder: "- a \<le> b \<longleftrightarrow> -b \<le> (a::extreal)"
-  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_le_minus)
-
-lemmas extreal_uminus_reorder =
-  extreal_uminus_eq_reorder extreal_uminus_less_reorder extreal_uminus_le_reorder
-
-lemma extreal_bot:
-  fixes x :: extreal assumes "\<And>B. x \<le> extreal B" shows "x = - \<infinity>"
-proof (cases x)
-  case (real r) with assms[of "r - 1"] show ?thesis by auto
-next case PInf with assms[of 0] show ?thesis by auto
-next case MInf then show ?thesis by simp
-qed
-
-lemma extreal_top:
-  fixes x :: extreal assumes "\<And>B. x \<ge> extreal B" shows "x = \<infinity>"
-proof (cases x)
-  case (real r) with assms[of "r + 1"] show ?thesis by auto
-next case MInf with assms[of 0] show ?thesis by auto
-next case PInf then show ?thesis by simp
-qed
-
-lemma
-  shows extreal_max[simp]: "extreal (max x y) = max (extreal x) (extreal y)"
-    and extreal_min[simp]: "extreal (min x y) = min (extreal x) (extreal y)"
-  by (simp_all add: min_def max_def)
-
-lemma extreal_max_0: "max 0 (extreal r) = extreal (max 0 r)"
-  by (auto simp: zero_extreal_def)
-
-lemma
-  fixes f :: "nat \<Rightarrow> extreal"
-  shows incseq_uminus[simp]: "incseq (\<lambda>x. - f x) \<longleftrightarrow> decseq f"
-  and decseq_uminus[simp]: "decseq (\<lambda>x. - f x) \<longleftrightarrow> incseq f"
-  unfolding decseq_def incseq_def by auto
-
-lemma incseq_extreal: "incseq f \<Longrightarrow> incseq (\<lambda>x. extreal (f x))"
-  unfolding incseq_def by auto
-
-lemma extreal_add_nonneg_nonneg:
-  fixes a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
-  using add_mono[of 0 a 0 b] by simp
-
-lemma image_eqD: "f ` A = B \<Longrightarrow> (\<forall>x\<in>A. f x \<in> B)"
-  by auto
-
-lemma incseq_setsumI:
-  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
-  assumes "\<And>i. 0 \<le> f i"
-  shows "incseq (\<lambda>i. setsum f {..< i})"
-proof (intro incseq_SucI)
-  fix n have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
-    using assms by (rule add_left_mono)
-  then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
-    by auto
-qed
-
-lemma incseq_setsumI2:
-  fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
-  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
-  shows "incseq (\<lambda>i. \<Sum>n\<in>A. f n i)"
-  using assms unfolding incseq_def by (auto intro: setsum_mono)
-
-subsubsection "Multiplication"
-
-instantiation extreal :: "{comm_monoid_mult, sgn}"
-begin
-
-definition "1 = extreal 1"
-
-function sgn_extreal where
-  "sgn (extreal r) = extreal (sgn r)"
-| "sgn \<infinity> = 1"
-| "sgn (-\<infinity>) = -1"
-by (auto intro: extreal_cases)
-termination proof qed (rule wf_empty)
-
-function times_extreal where
-"extreal r * extreal p = extreal (r * p)" |
-"extreal r * \<infinity> = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
-"\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
-"extreal r * -\<infinity> = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
-"-\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
-"\<infinity> * \<infinity> = \<infinity>" |
-"-\<infinity> * \<infinity> = -\<infinity>" |
-"\<infinity> * -\<infinity> = -\<infinity>" |
-"-\<infinity> * -\<infinity> = \<infinity>"
-proof -
-  case (goal1 P x)
-  moreover then obtain a b where "x = (a, b)" by (cases x) auto
-  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
-qed simp_all
-termination by (relation "{}") simp
-
-instance
-proof
-  fix a :: extreal show "1 * a = a"
-    by (cases a) (simp_all add: one_extreal_def)
-  fix b :: extreal show "a * b = b * a"
-    by (cases rule: extreal2_cases[of a b]) simp_all
-  fix c :: extreal show "a * b * c = a * (b * c)"
-    by (cases rule: extreal3_cases[of a b c])
-       (simp_all add: zero_extreal_def zero_less_mult_iff)
-qed
-end
-
-lemma real_of_extreal_le_1:
-  fixes a :: extreal shows "a \<le> 1 \<Longrightarrow> real a \<le> 1"
-  by (cases a) (auto simp: one_extreal_def)
-
-lemma abs_extreal_one[simp]: "\<bar>1\<bar> = (1::extreal)"
-  unfolding one_extreal_def by simp
-
-lemma extreal_mult_zero[simp]:
-  fixes a :: extreal shows "a * 0 = 0"
-  by (cases a) (simp_all add: zero_extreal_def)
-
-lemma extreal_zero_mult[simp]:
-  fixes a :: extreal shows "0 * a = 0"
-  by (cases a) (simp_all add: zero_extreal_def)
-
-lemma extreal_m1_less_0[simp]:
-  "-(1::extreal) < 0"
-  by (simp add: zero_extreal_def one_extreal_def)
-
-lemma extreal_zero_m1[simp]:
-  "1 \<noteq> (0::extreal)"
-  by (simp add: zero_extreal_def one_extreal_def)
-
-lemma extreal_times_0[simp]:
-  fixes x :: extreal shows "0 * x = 0"
-  by (cases x) (auto simp: zero_extreal_def)
-
-lemma extreal_times[simp]:
-  "1 \<noteq> \<infinity>" "\<infinity> \<noteq> 1"
-  "1 \<noteq> -\<infinity>" "-\<infinity> \<noteq> 1"
-  by (auto simp add: times_extreal_def one_extreal_def)
-
-lemma extreal_plus_1[simp]:
-  "1 + extreal r = extreal (r + 1)" "extreal r + 1 = extreal (r + 1)"
-  "1 + -\<infinity> = -\<infinity>" "-\<infinity> + 1 = -\<infinity>"
-  unfolding one_extreal_def by auto
-
-lemma extreal_zero_times[simp]:
-  fixes a b :: extreal shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_mult_eq_PInfty[simp]:
-  shows "a * b = \<infinity> \<longleftrightarrow>
-    (a = \<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = -\<infinity>)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_mult_eq_MInfty[simp]:
-  shows "a * b = -\<infinity> \<longleftrightarrow>
-    (a = \<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = -\<infinity>)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_0_less_1[simp]: "0 < (1::extreal)"
-  by (simp_all add: zero_extreal_def one_extreal_def)
-
-lemma extreal_zero_one[simp]: "0 \<noteq> (1::extreal)"
-  by (simp_all add: zero_extreal_def one_extreal_def)
-
-lemma extreal_mult_minus_left[simp]:
-  fixes a b :: extreal shows "-a * b = - (a * b)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_mult_minus_right[simp]:
-  fixes a b :: extreal shows "a * -b = - (a * b)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_mult_infty[simp]:
-  "a * \<infinity> = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
-  by (cases a) auto
-
-lemma extreal_infty_mult[simp]:
-  "\<infinity> * a = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
-  by (cases a) auto
-
-lemma extreal_mult_strict_right_mono:
-  assumes "a < b" and "0 < c" "c < \<infinity>"
-  shows "a * c < b * c"
-  using assms
-  by (cases rule: extreal3_cases[of a b c])
-     (auto simp: zero_le_mult_iff extreal_less_PInfty)
-
-lemma extreal_mult_strict_left_mono:
-  "\<lbrakk> a < b ; 0 < c ; c < \<infinity>\<rbrakk> \<Longrightarrow> c * a < c * b"
-  using extreal_mult_strict_right_mono by (simp add: mult_commute[of c])
-
-lemma extreal_mult_right_mono:
-  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a*c \<le> b*c"
-  using assms
-  apply (cases "c = 0") apply simp
-  by (cases rule: extreal3_cases[of a b c])
-     (auto simp: zero_le_mult_iff extreal_less_PInfty)
-
-lemma extreal_mult_left_mono:
-  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> c * a \<le> c * b"
-  using extreal_mult_right_mono by (simp add: mult_commute[of c])
-
-lemma zero_less_one_extreal[simp]: "0 \<le> (1::extreal)"
-  by (simp add: one_extreal_def zero_extreal_def)
-
-lemma extreal_0_le_mult[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * (b :: extreal)"
-  by (cases rule: extreal2_cases[of a b]) (auto simp: mult_nonneg_nonneg)
-
-lemma extreal_right_distrib:
-  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> r * (a + b) = r * a + r * b"
-  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
-
-lemma extreal_left_distrib:
-  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> (a + b) * r = a * r + b * r"
-  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
-
-lemma extreal_mult_le_0_iff:
-  fixes a b :: extreal
-  shows "a * b \<le> 0 \<longleftrightarrow> (0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b)"
-  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_le_0_iff)
-
-lemma extreal_zero_le_0_iff:
-  fixes a b :: extreal
-  shows "0 \<le> a * b \<longleftrightarrow> (0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0)"
-  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_le_mult_iff)
-
-lemma extreal_mult_less_0_iff:
-  fixes a b :: extreal
-  shows "a * b < 0 \<longleftrightarrow> (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b)"
-  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_less_0_iff)
-
-lemma extreal_zero_less_0_iff:
-  fixes a b :: extreal
-  shows "0 < a * b \<longleftrightarrow> (0 < a \<and> 0 < b) \<or> (a < 0 \<and> b < 0)"
-  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_less_mult_iff)
-
-lemma extreal_distrib:
-  fixes a b c :: extreal
-  assumes "a \<noteq> \<infinity> \<or> b \<noteq> -\<infinity>" "a \<noteq> -\<infinity> \<or> b \<noteq> \<infinity>" "\<bar>c\<bar> \<noteq> \<infinity>"
-  shows "(a + b) * c = a * c + b * c"
-  using assms
-  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
-
-lemma extreal_le_epsilon:
-  fixes x y :: extreal
-  assumes "ALL e. 0 < e --> x <= y + e"
-  shows "x <= y"
-proof-
-{ assume a: "EX r. y = extreal r"
-  from this obtain r where r_def: "y = extreal r" by auto
-  { assume "x=(-\<infinity>)" hence ?thesis by auto }
-  moreover
-  { assume "~(x=(-\<infinity>))"
-    from this obtain p where p_def: "x = extreal p"
-    using a assms[rule_format, of 1] by (cases x) auto
-    { fix e have "0 < e --> p <= r + e"
-      using assms[rule_format, of "extreal e"] p_def r_def by auto }
-    hence "p <= r" apply (subst field_le_epsilon) by auto
-    hence ?thesis using r_def p_def by auto
-  } ultimately have ?thesis by blast
-}
-moreover
-{ assume "y=(-\<infinity>) | y=\<infinity>" hence ?thesis
-    using assms[rule_format, of 1] by (cases x) auto
-} ultimately show ?thesis by (cases y) auto
-qed
-
-
-lemma extreal_le_epsilon2:
-  fixes x y :: extreal
-  assumes "ALL e. 0 < e --> x <= y + extreal e"
-  shows "x <= y"
-proof-
-{ fix e :: extreal assume "e>0"
-  { assume "e=\<infinity>" hence "x<=y+e" by auto }
-  moreover
-  { assume "e~=\<infinity>"
-    from this obtain r where "e = extreal r" using `e>0` apply (cases e) by auto
-    hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto
-  } ultimately have "x<=y+e" by blast
-} from this show ?thesis using extreal_le_epsilon by auto
-qed
-
-lemma extreal_le_real:
-  fixes x y :: extreal
-  assumes "ALL z. x <= extreal z --> y <= extreal z"
-  shows "y <= x"
-by (metis assms extreal.exhaust extreal_bot extreal_less_eq(1)
-          extreal_less_eq(2) order_refl uminus_extreal.simps(2))
-
-lemma extreal_le_extreal:
-  fixes x y :: extreal
-  assumes "\<And>B. B < x \<Longrightarrow> B <= y"
-  shows "x <= y"
-by (metis assms extreal_dense leD linorder_le_less_linear)
-
-lemma extreal_ge_extreal:
-  fixes x y :: extreal
-  assumes "ALL B. B>x --> B >= y"
-  shows "x >= y"
-by (metis assms extreal_dense leD linorder_le_less_linear)
-
-lemma setprod_extreal_0:
-  fixes f :: "'a \<Rightarrow> extreal"
-  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
-proof cases
-  assume "finite A"
-  then show ?thesis by (induct A) auto
-qed auto
-
-lemma setprod_extreal_pos:
-  fixes f :: "'a \<Rightarrow> extreal" assumes pos: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i" shows "0 \<le> (\<Prod>i\<in>I. f i)"
-proof cases
-  assume "finite I" from this pos show ?thesis by induct auto
-qed simp
-
-lemma setprod_PInf:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
-  shows "(\<Prod>i\<in>I. f i) = \<infinity> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<infinity>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
-proof cases
-  assume "finite I" from this assms show ?thesis
-  proof (induct I)
-    case (insert i I)
-    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_extreal_pos)
-    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
-    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
-      using setprod_extreal_pos[of I f] pos
-      by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto
-    also have "\<dots> \<longleftrightarrow> finite (insert i I) \<and> (\<exists>j\<in>insert i I. f j = \<infinity>) \<and> (\<forall>j\<in>insert i I. f j \<noteq> 0)"
-      using insert by (auto simp: setprod_extreal_0)
-    finally show ?case .
-  qed simp
-qed simp
-
-lemma setprod_extreal: "(\<Prod>i\<in>A. extreal (f i)) = extreal (setprod f A)"
-proof cases
-  assume "finite A" then show ?thesis
-    by induct (auto simp: one_extreal_def)
-qed (simp add: one_extreal_def)
-
-subsubsection {* Power *}
-
-instantiation extreal :: power
-begin
-primrec power_extreal where
-  "power_extreal x 0 = 1" |
-  "power_extreal x (Suc n) = x * x ^ n"
-instance ..
-end
-
-lemma extreal_power[simp]: "(extreal x) ^ n = extreal (x^n)"
-  by (induct n) (auto simp: one_extreal_def)
-
-lemma extreal_power_PInf[simp]: "\<infinity> ^ n = (if n = 0 then 1 else \<infinity>)"
-  by (induct n) (auto simp: one_extreal_def)
-
-lemma extreal_power_uminus[simp]:
-  fixes x :: extreal
-  shows "(- x) ^ n = (if even n then x ^ n else - (x^n))"
-  by (induct n) (auto simp: one_extreal_def)
-
-lemma extreal_power_number_of[simp]:
-  "(number_of num :: extreal) ^ n = extreal (number_of num ^ n)"
-  by (induct n) (auto simp: one_extreal_def)
-
-lemma zero_le_power_extreal[simp]:
-  fixes a :: extreal assumes "0 \<le> a"
-  shows "0 \<le> a ^ n"
-  using assms by (induct n) (auto simp: extreal_zero_le_0_iff)
-
-subsubsection {* Subtraction *}
-
-lemma extreal_minus_minus_image[simp]:
-  fixes S :: "extreal set"
-  shows "uminus ` uminus ` S = S"
-  by (auto simp: image_iff)
-
-lemma extreal_uminus_lessThan[simp]:
-  fixes a :: extreal shows "uminus ` {..<a} = {-a<..}"
-proof (safe intro!: image_eqI)
-  fix x assume "-a < x"
-  then have "- x < - (- a)" by (simp del: extreal_uminus_uminus)
-  then show "- x < a" by simp
-qed auto
-
-lemma extreal_uminus_greaterThan[simp]:
-  "uminus ` {(a::extreal)<..} = {..<-a}"
-  by (metis extreal_uminus_lessThan extreal_uminus_uminus
-            extreal_minus_minus_image)
-
-instantiation extreal :: minus
-begin
-definition "x - y = x + -(y::extreal)"
-instance ..
-end
-
-lemma extreal_minus[simp]:
-  "extreal r - extreal p = extreal (r - p)"
-  "-\<infinity> - extreal r = -\<infinity>"
-  "extreal r - \<infinity> = -\<infinity>"
-  "\<infinity> - x = \<infinity>"
-  "-\<infinity> - \<infinity> = -\<infinity>"
-  "x - -y = x + y"
-  "x - 0 = x"
-  "0 - x = -x"
-  by (simp_all add: minus_extreal_def)
-
-lemma extreal_x_minus_x[simp]:
-  "x - x = (if \<bar>x\<bar> = \<infinity> then \<infinity> else 0)"
-  by (cases x) simp_all
-
-lemma extreal_eq_minus_iff:
-  fixes x y z :: extreal
-  shows "x = z - y \<longleftrightarrow>
-    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y = z) \<and>
-    (y = -\<infinity> \<longrightarrow> x = \<infinity>) \<and>
-    (y = \<infinity> \<longrightarrow> z = \<infinity> \<longrightarrow> x = \<infinity>) \<and>
-    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>)"
-  by (cases rule: extreal3_cases[of x y z]) auto
-
-lemma extreal_eq_minus:
-  fixes x y z :: extreal
-  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x = z - y \<longleftrightarrow> x + y = z"
-  by (auto simp: extreal_eq_minus_iff)
-
-lemma extreal_less_minus_iff:
-  fixes x y z :: extreal
-  shows "x < z - y \<longleftrightarrow>
-    (y = \<infinity> \<longrightarrow> z = \<infinity> \<and> x \<noteq> \<infinity>) \<and>
-    (y = -\<infinity> \<longrightarrow> x \<noteq> \<infinity>) \<and>
-    (\<bar>y\<bar> \<noteq> \<infinity>\<longrightarrow> x + y < z)"
-  by (cases rule: extreal3_cases[of x y z]) auto
-
-lemma extreal_less_minus:
-  fixes x y z :: extreal
-  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x < z - y \<longleftrightarrow> x + y < z"
-  by (auto simp: extreal_less_minus_iff)
-
-lemma extreal_le_minus_iff:
-  fixes x y z :: extreal
-  shows "x \<le> z - y \<longleftrightarrow>
-    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>) \<and>
-    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y \<le> z)"
-  by (cases rule: extreal3_cases[of x y z]) auto
-
-lemma extreal_le_minus:
-  fixes x y z :: extreal
-  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x \<le> z - y \<longleftrightarrow> x + y \<le> z"
-  by (auto simp: extreal_le_minus_iff)
-
-lemma extreal_minus_less_iff:
-  fixes x y z :: extreal
-  shows "x - y < z \<longleftrightarrow>
-    y \<noteq> -\<infinity> \<and> (y = \<infinity> \<longrightarrow> x \<noteq> \<infinity> \<and> z \<noteq> -\<infinity>) \<and>
-    (y \<noteq> \<infinity> \<longrightarrow> x < z + y)"
-  by (cases rule: extreal3_cases[of x y z]) auto
-
-lemma extreal_minus_less:
-  fixes x y z :: extreal
-  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y < z \<longleftrightarrow> x < z + y"
-  by (auto simp: extreal_minus_less_iff)
-
-lemma extreal_minus_le_iff:
-  fixes x y z :: extreal
-  shows "x - y \<le> z \<longleftrightarrow>
-    (y = -\<infinity> \<longrightarrow> z = \<infinity>) \<and>
-    (y = \<infinity> \<longrightarrow> x = \<infinity> \<longrightarrow> z = \<infinity>) \<and>
-    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x \<le> z + y)"
-  by (cases rule: extreal3_cases[of x y z]) auto
-
-lemma extreal_minus_le:
-  fixes x y z :: extreal
-  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y \<le> z \<longleftrightarrow> x \<le> z + y"
-  by (auto simp: extreal_minus_le_iff)
-
-lemma extreal_minus_eq_minus_iff:
-  fixes a b c :: extreal
-  shows "a - b = a - c \<longleftrightarrow>
-    b = c \<or> a = \<infinity> \<or> (a = -\<infinity> \<and> b \<noteq> -\<infinity> \<and> c \<noteq> -\<infinity>)"
-  by (cases rule: extreal3_cases[of a b c]) auto
-
-lemma extreal_add_le_add_iff:
-  "c + a \<le> c + b \<longleftrightarrow>
-    a \<le> b \<or> c = \<infinity> \<or> (c = -\<infinity> \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>)"
-  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
-
-lemma extreal_mult_le_mult_iff:
-  "\<bar>c\<bar> \<noteq> \<infinity> \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
-  by (cases rule: extreal3_cases[of a b c]) (simp_all add: mult_le_cancel_left)
-
-lemma extreal_minus_mono:
-  fixes A B C D :: extreal assumes "A \<le> B" "D \<le> C"
-  shows "A - C \<le> B - D"
-  using assms
-  by (cases rule: extreal3_cases[case_product extreal_cases, of A B C D]) simp_all
-
-lemma real_of_extreal_minus:
-  "real (a - b) = (if \<bar>a\<bar> = \<infinity> \<or> \<bar>b\<bar> = \<infinity> then 0 else real a - real b)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_diff_positive:
-  fixes a b :: extreal shows "a \<le> b \<Longrightarrow> 0 \<le> b - a"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_between:
-  fixes x e :: extreal
-  assumes "\<bar>x\<bar> \<noteq> \<infinity>" "0 < e"
-  shows "x - e < x" "x < x + e"
-using assms apply (cases x, cases e) apply auto
-using assms by (cases x, cases e) auto
-
-subsubsection {* Division *}
-
-instantiation extreal :: inverse
-begin
-
-function inverse_extreal where
-"inverse (extreal r) = (if r = 0 then \<infinity> else extreal (inverse r))" |
-"inverse \<infinity> = 0" |
-"inverse (-\<infinity>) = 0"
-  by (auto intro: extreal_cases)
-termination by (relation "{}") simp
-
-definition "x / y = x * inverse (y :: extreal)"
-
-instance proof qed
-end
-
-lemma real_of_extreal_inverse[simp]:
-  fixes a :: extreal
-  shows "real (inverse a) = 1 / real a"
-  by (cases a) (auto simp: inverse_eq_divide)
-
-lemma extreal_inverse[simp]:
-  "inverse 0 = \<infinity>"
-  "inverse (1::extreal) = 1"
-  by (simp_all add: one_extreal_def zero_extreal_def)
-
-lemma extreal_divide[simp]:
-  "extreal r / extreal p = (if p = 0 then extreal r * \<infinity> else extreal (r / p))"
-  unfolding divide_extreal_def by (auto simp: divide_real_def)
-
-lemma extreal_divide_same[simp]:
-  "x / x = (if \<bar>x\<bar> = \<infinity> \<or> x = 0 then 0 else 1)"
-  by (cases x)
-     (simp_all add: divide_real_def divide_extreal_def one_extreal_def)
-
-lemma extreal_inv_inv[simp]:
-  "inverse (inverse x) = (if x \<noteq> -\<infinity> then x else \<infinity>)"
-  by (cases x) auto
-
-lemma extreal_inverse_minus[simp]:
-  "inverse (- x) = (if x = 0 then \<infinity> else -inverse x)"
-  by (cases x) simp_all
-
-lemma extreal_uminus_divide[simp]:
-  fixes x y :: extreal shows "- x / y = - (x / y)"
-  unfolding divide_extreal_def by simp
-
-lemma extreal_divide_Infty[simp]:
-  "x / \<infinity> = 0" "x / -\<infinity> = 0"
-  unfolding divide_extreal_def by simp_all
-
-lemma extreal_divide_one[simp]:
-  "x / 1 = (x::extreal)"
-  unfolding divide_extreal_def by simp
-
-lemma extreal_divide_extreal[simp]:
-  "\<infinity> / extreal r = (if 0 \<le> r then \<infinity> else -\<infinity>)"
-  unfolding divide_extreal_def by simp
-
-lemma zero_le_divide_extreal[simp]:
-  fixes a :: extreal assumes "0 \<le> a" "0 \<le> b"
-  shows "0 \<le> a / b"
-  using assms by (cases rule: extreal2_cases[of a b]) (auto simp: zero_le_divide_iff)
-
-lemma extreal_le_divide_pos:
-  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> x * y \<le> z"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_divide_le_pos:
-  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> z \<le> x * y"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_le_divide_neg:
-  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> z \<le> x * y"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_divide_le_neg:
-  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> x * y \<le> z"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_inverse_antimono_strict:
-  fixes x y :: extreal
-  shows "0 \<le> x \<Longrightarrow> x < y \<Longrightarrow> inverse y < inverse x"
-  by (cases rule: extreal2_cases[of x y]) auto
-
-lemma extreal_inverse_antimono:
-  fixes x y :: extreal
-  shows "0 \<le> x \<Longrightarrow> x <= y \<Longrightarrow> inverse y <= inverse x"
-  by (cases rule: extreal2_cases[of x y]) auto
-
-lemma inverse_inverse_Pinfty_iff[simp]:
-  "inverse x = \<infinity> \<longleftrightarrow> x = 0"
-  by (cases x) auto
-
-lemma extreal_inverse_eq_0:
-  "inverse x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity>"
-  by (cases x) auto
-
-lemma extreal_0_gt_inverse:
-  fixes x :: extreal shows "0 < inverse x \<longleftrightarrow> x \<noteq> \<infinity> \<and> 0 \<le> x"
-  by (cases x) auto
-
-lemma extreal_mult_less_right:
-  assumes "b * a < c * a" "0 < a" "a < \<infinity>"
-  shows "b < c"
-  using assms
-  by (cases rule: extreal3_cases[of a b c])
-     (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
-
-lemma extreal_power_divide:
-  "y \<noteq> 0 \<Longrightarrow> (x / y :: extreal) ^ n = x^n / y^n"
-  by (cases rule: extreal2_cases[of x y])
-     (auto simp: one_extreal_def zero_extreal_def power_divide not_le
-                 power_less_zero_eq zero_le_power_iff)
-
-lemma extreal_le_mult_one_interval:
-  fixes x y :: extreal
-  assumes y: "y \<noteq> -\<infinity>"
-  assumes z: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
-  shows "x \<le> y"
-proof (cases x)
-  case PInf with z[of "1 / 2"] show "x \<le> y" by (simp add: one_extreal_def)
-next
-  case (real r) note r = this
-  show "x \<le> y"
-  proof (cases y)
-    case (real p) note p = this
-    have "r \<le> p"
-    proof (rule field_le_mult_one_interval)
-      fix z :: real assume "0 < z" and "z < 1"
-      with z[of "extreal z"]
-      show "z * r \<le> p" using p r by (auto simp: zero_le_mult_iff one_extreal_def)
-    qed
-    then show "x \<le> y" using p r by simp
-  qed (insert y, simp_all)
-qed simp
-
-subsection "Complete lattice"
-
-instantiation extreal :: lattice
-begin
-definition [simp]: "sup x y = (max x y :: extreal)"
-definition [simp]: "inf x y = (min x y :: extreal)"
-instance proof qed simp_all
-end
-
-instantiation extreal :: complete_lattice
-begin
-
-definition "bot = -\<infinity>"
-definition "top = \<infinity>"
-
-definition "Sup S = (LEAST z. ALL x:S. x <= z :: extreal)"
-definition "Inf S = (GREATEST z. ALL x:S. z <= x :: extreal)"
-
-lemma extreal_complete_Sup:
-  fixes S :: "extreal set" assumes "S \<noteq> {}"
-  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
-proof cases
-  assume "\<exists>x. \<forall>a\<in>S. a \<le> extreal x"
-  then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> extreal y" by auto
-  then have "\<infinity> \<notin> S" by force
-  show ?thesis
-  proof cases
-    assume "S = {-\<infinity>}"
-    then show ?thesis by (auto intro!: exI[of _ "-\<infinity>"])
-  next
-    assume "S \<noteq> {-\<infinity>}"
-    with `S \<noteq> {}` `\<infinity> \<notin> S` obtain x where "x \<in> S - {-\<infinity>}" "x \<noteq> \<infinity>" by auto
-    with y `\<infinity> \<notin> S` have "\<forall>z\<in>real ` (S - {-\<infinity>}). z \<le> y"
-      by (auto simp: real_of_extreal_ord_simps)
-    with reals_complete2[of "real ` (S - {-\<infinity>})"] `x \<in> S - {-\<infinity>}`
-    obtain s where s:
-       "\<forall>y\<in>S - {-\<infinity>}. real y \<le> s" "\<And>z. (\<forall>y\<in>S - {-\<infinity>}. real y \<le> z) \<Longrightarrow> s \<le> z"
-       by auto
-    show ?thesis
-    proof (safe intro!: exI[of _ "extreal s"])
-      fix z assume "z \<in> S" with `\<infinity> \<notin> S` show "z \<le> extreal s"
-      proof (cases z)
-        case (real r)
-        then show ?thesis
-          using s(1)[rule_format, of z] `z \<in> S` `z = extreal r` by auto
-      qed auto
-    next
-      fix z assume *: "\<forall>y\<in>S. y \<le> z"
-      with `S \<noteq> {-\<infinity>}` `S \<noteq> {}` show "extreal s \<le> z"
-      proof (cases z)
-        case (real u)
-        with * have "s \<le> u"
-          by (intro s(2)[of u]) (auto simp: real_of_extreal_ord_simps)
-        then show ?thesis using real by simp
-      qed auto
-    qed
-  qed
-next
-  assume *: "\<not> (\<exists>x. \<forall>a\<in>S. a \<le> extreal x)"
-  show ?thesis
-  proof (safe intro!: exI[of _ \<infinity>])
-    fix y assume **: "\<forall>z\<in>S. z \<le> y"
-    with * show "\<infinity> \<le> y"
-    proof (cases y)
-      case MInf with * ** show ?thesis by (force simp: not_le)
-    qed auto
-  qed simp
-qed
-
-lemma extreal_complete_Inf:
-  fixes S :: "extreal set" assumes "S ~= {}"
-  shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)"
-proof-
-def S1 == "uminus ` S"
-hence "S1 ~= {}" using assms by auto
-from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)"
-   using extreal_complete_Sup[of S1] by auto
-{ fix z assume "ALL y:S. z <= y"
-  hence "ALL y:S1. y <= -z" unfolding S1_def by auto
-  hence "x <= -z" using x_def by auto
-  hence "z <= -x"
-    apply (subst extreal_uminus_uminus[symmetric])
-    unfolding extreal_minus_le_minus . }
-moreover have "(ALL y:S. -x <= y)"
-   using x_def unfolding S1_def
-   apply simp
-   apply (subst (3) extreal_uminus_uminus[symmetric])
-   unfolding extreal_minus_le_minus by simp
-ultimately show ?thesis by auto
-qed
-
-lemma extreal_complete_uminus_eq:
-  fixes S :: "extreal set"
-  shows "(\<forall>y\<in>uminus`S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>uminus`S. y \<le> z) \<longrightarrow> x \<le> z)
-     \<longleftrightarrow> (\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
-  by simp (metis extreal_minus_le_minus extreal_uminus_uminus)
-
-lemma extreal_Sup_uminus_image_eq:
-  fixes S :: "extreal set"
-  shows "Sup (uminus ` S) = - Inf S"
-proof cases
-  assume "S = {}"
-  moreover have "(THE x. All (op \<le> x)) = (-\<infinity>::extreal)"
-    by (rule the_equality) (auto intro!: extreal_bot)
-  moreover have "(SOME x. \<forall>y. y \<le> x) = (\<infinity>::extreal)"
-    by (rule some_equality) (auto intro!: extreal_top)
-  ultimately show ?thesis unfolding Inf_extreal_def Sup_extreal_def
-    Least_def Greatest_def GreatestM_def by simp
-next
-  assume "S \<noteq> {}"
-  with extreal_complete_Sup[of "uminus`S"]
-  obtain x where x: "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
-    unfolding extreal_complete_uminus_eq by auto
-  show "Sup (uminus ` S) = - Inf S"
-    unfolding Inf_extreal_def Greatest_def GreatestM_def
-  proof (intro someI2[of _ _ "\<lambda>x. Sup (uminus`S) = - x"])
-    show "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> -x)"
-      using x .
-    fix x' assume "(\<forall>y\<in>S. x' \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> x')"
-    then have "(\<forall>y\<in>uminus`S. y \<le> - x') \<and> (\<forall>y. (\<forall>z\<in>uminus`S. z \<le> y) \<longrightarrow> - x' \<le> y)"
-      unfolding extreal_complete_uminus_eq by simp
-    then show "Sup (uminus ` S) = -x'"
-      unfolding Sup_extreal_def extreal_uminus_eq_iff
-      by (intro Least_equality) auto
-  qed
-qed
-
-instance
-proof
-  { fix x :: extreal and A
-    show "bot <= x" by (cases x) (simp_all add: bot_extreal_def)
-    show "x <= top" by (simp add: top_extreal_def) }
-
-  { fix x :: extreal and A assume "x : A"
-    with extreal_complete_Sup[of A]
-    obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
-    hence "x <= s" using `x : A` by auto
-    also have "... = Sup A" using s unfolding Sup_extreal_def
-      by (auto intro!: Least_equality[symmetric])
-    finally show "x <= Sup A" . }
-  note le_Sup = this
-
-  { fix x :: extreal and A assume *: "!!z. (z : A ==> z <= x)"
-    show "Sup A <= x"
-    proof (cases "A = {}")
-      case True
-      hence "Sup A = -\<infinity>" unfolding Sup_extreal_def
-        by (auto intro!: Least_equality)
-      thus "Sup A <= x" by simp
-    next
-      case False
-      with extreal_complete_Sup[of A]
-      obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
-      hence "Sup A = s"
-        unfolding Sup_extreal_def by (auto intro!: Least_equality)
-      also have "s <= x" using * s by auto
-      finally show "Sup A <= x" .
-    qed }
-  note Sup_le = this
-
-  { fix x :: extreal and A assume "x \<in> A"
-    with le_Sup[of "-x" "uminus`A"] show "Inf A \<le> x"
-      unfolding extreal_Sup_uminus_image_eq by simp }
-
-  { fix x :: extreal and A assume *: "!!z. (z : A ==> x <= z)"
-    with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
-      unfolding extreal_Sup_uminus_image_eq by force }
-qed
-end
-
-lemma extreal_SUPR_uminus:
-  fixes f :: "'a => extreal"
-  shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
-  unfolding SUPR_def INFI_def
-  using extreal_Sup_uminus_image_eq[of "f`R"]
-  by (simp add: image_image)
-
-lemma extreal_INFI_uminus:
-  fixes f :: "'a => extreal"
-  shows "(INF i : R. -(f i)) = -(SUP i : R. f i)"
-  using extreal_SUPR_uminus[of _ "\<lambda>x. - f x"] by simp
-
-lemma extreal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::extreal set)"
-  using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image)
-
-lemma extreal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: extreal set)"
-  by (auto intro!: inj_onI)
-
-lemma extreal_image_uminus_shift:
-  fixes X Y :: "extreal set" shows "uminus ` X = Y \<longleftrightarrow> X = uminus ` Y"
-proof
-  assume "uminus ` X = Y"
-  then have "uminus ` uminus ` X = uminus ` Y"
-    by (simp add: inj_image_eq_iff)
-  then show "X = uminus ` Y" by (simp add: image_image)
-qed (simp add: image_image)
-
-lemma Inf_extreal_iff:
-  fixes z :: extreal
-  shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x<y) <-> Inf X < y"
-  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
-            order_less_le_trans)
-
-lemma Sup_eq_MInfty:
-  fixes S :: "extreal set" shows "Sup S = -\<infinity> \<longleftrightarrow> S = {} \<or> S = {-\<infinity>}"
-proof
-  assume a: "Sup S = -\<infinity>"
-  with complete_lattice_class.Sup_upper[of _ S]
-  show "S={} \<or> S={-\<infinity>}" by auto
-next
-  assume "S={} \<or> S={-\<infinity>}" then show "Sup S = -\<infinity>"
-    unfolding Sup_extreal_def by (auto intro!: Least_equality)
-qed
-
-lemma Inf_eq_PInfty:
-  fixes S :: "extreal set" shows "Inf S = \<infinity> \<longleftrightarrow> S = {} \<or> S = {\<infinity>}"
-  using Sup_eq_MInfty[of "uminus`S"]
-  unfolding extreal_Sup_uminus_image_eq extreal_image_uminus_shift by simp
-
-lemma Inf_eq_MInfty: "-\<infinity> : S ==> Inf S = -\<infinity>"
-  unfolding Inf_extreal_def
-  by (auto intro!: Greatest_equality)
-
-lemma Sup_eq_PInfty: "\<infinity> : S ==> Sup S = \<infinity>"
-  unfolding Sup_extreal_def
-  by (auto intro!: Least_equality)
-
-lemma extreal_SUPI:
-  fixes x :: extreal
-  assumes "!!i. i : A ==> f i <= x"
-  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
-  shows "(SUP i:A. f i) = x"
-  unfolding SUPR_def Sup_extreal_def
-  using assms by (auto intro!: Least_equality)
-
-lemma extreal_INFI:
-  fixes x :: extreal
-  assumes "!!i. i : A ==> f i >= x"
-  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
-  shows "(INF i:A. f i) = x"
-  unfolding INFI_def Inf_extreal_def
-  using assms by (auto intro!: Greatest_equality)
-
-lemma Sup_extreal_close:
-  fixes e :: extreal
-  assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
-  shows "\<exists>x\<in>S. Sup S - e < x"
-  using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1])
-
-lemma Inf_extreal_close:
-  fixes e :: extreal assumes "\<bar>Inf X\<bar> \<noteq> \<infinity>" "0 < e"
-  shows "\<exists>x\<in>X. x < Inf X + e"
-proof (rule Inf_less_iff[THEN iffD1])
-  show "Inf X < Inf X + e" using assms
-    by (cases e) auto
-qed
-
-lemma Sup_eq_top_iff:
-  fixes A :: "'a::{complete_lattice, linorder} set"
-  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
-proof
-  assume *: "Sup A = top"
-  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
-  proof (intro allI impI)
-    fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
-      unfolding less_Sup_iff by auto
-  qed
-next
-  assume *: "\<forall>x<top. \<exists>i\<in>A. x < i"
-  show "Sup A = top"
-  proof (rule ccontr)
-    assume "Sup A \<noteq> top"
-    with top_greatest[of "Sup A"]
-    have "Sup A < top" unfolding le_less by auto
-    then have "Sup A < Sup A"
-      using * unfolding less_Sup_iff by auto
-    then show False by auto
-  qed
-qed
-
-lemma SUP_eq_top_iff:
-  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
-  shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
-  unfolding SUPR_def Sup_eq_top_iff by auto
-
-lemma SUP_nat_Infty: "(SUP i::nat. extreal (real i)) = \<infinity>"
-proof -
-  { fix x assume "x \<noteq> \<infinity>"
-    then have "\<exists>k::nat. x < extreal (real k)"
-    proof (cases x)
-      case MInf then show ?thesis by (intro exI[of _ 0]) auto
-    next
-      case (real r)
-      moreover obtain k :: nat where "r < real k"
-        using ex_less_of_nat by (auto simp: real_eq_of_nat)
-      ultimately show ?thesis by auto
-    qed simp }
-  then show ?thesis
-    using SUP_eq_top_iff[of UNIV "\<lambda>n::nat. extreal (real n)"]
-    by (auto simp: top_extreal_def)
-qed
-
-lemma extreal_le_Sup:
-  fixes x :: extreal
-  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
-(is "?lhs <-> ?rhs")
-proof-
-{ assume "?rhs"
-  { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
-    from this obtain y where y_def: "(SUP i:A. f i)<y & y<x" using extreal_dense by auto
-    from this obtain i where "i : A & y <= f i" using `?rhs` by auto
-    hence "y <= (SUP i:A. f i)" using le_SUPI[of i A f] by auto
-    hence False using y_def by auto
-  } hence "?lhs" by auto
-}
-moreover
-{ assume "?lhs" hence "?rhs"
-  by (metis Collect_def Collect_mem_eq SUP_leI assms atLeastatMost_empty atLeastatMost_empty_iff
-      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
-} ultimately show ?thesis by auto
-qed
-
-lemma extreal_Inf_le:
-  fixes x :: extreal
-  shows "((INF i:A. f i) <= x) <-> (ALL y. x < y --> (EX i. i : A & f i <= y))"
-(is "?lhs <-> ?rhs")
-proof-
-{ assume "?rhs"
-  { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le)
-    from this obtain y where y_def: "x<y & y<(INF i:A. f i)" using extreal_dense by auto
-    from this obtain i where "i : A & f i <= y" using `?rhs` by auto
-    hence "(INF i:A. f i) <= y" using INF_leI[of i A f] by auto
-    hence False using y_def by auto
-  } hence "?lhs" by auto
-}
-moreover
-{ assume "?lhs" hence "?rhs"
-  by (metis Collect_def Collect_mem_eq le_INFI assms atLeastatMost_empty atLeastatMost_empty_iff
-      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
-} ultimately show ?thesis by auto
-qed
-
-lemma Inf_less:
-  fixes x :: extreal
-  assumes "(INF i:A. f i) < x"
-  shows "EX i. i : A & f i <= x"
-proof(rule ccontr)
-  assume "~ (EX i. i : A & f i <= x)"
-  hence "ALL i:A. f i > x" by auto
-  hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto
-  thus False using assms by auto
-qed
-
-lemma same_INF:
-  assumes "ALL e:A. f e = g e"
-  shows "(INF e:A. f e) = (INF e:A. g e)"
-proof-
-have "f ` A = g ` A" unfolding image_def using assms by auto
-thus ?thesis unfolding INFI_def by auto
-qed
-
-lemma same_SUP:
-  assumes "ALL e:A. f e = g e"
-  shows "(SUP e:A. f e) = (SUP e:A. g e)"
-proof-
-have "f ` A = g ` A" unfolding image_def using assms by auto
-thus ?thesis unfolding SUPR_def by auto
-qed
-
-lemma SUPR_eq:
-  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<le> g j"
-  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<le> f i"
-  shows "(SUP i:A. f i) = (SUP j:B. g j)"
-proof (intro antisym)
-  show "(SUP i:A. f i) \<le> (SUP j:B. g j)"
-    using assms by (metis SUP_leI le_SUPI2)
-  show "(SUP i:B. g i) \<le> (SUP j:A. f j)"
-    using assms by (metis SUP_leI le_SUPI2)
-qed
-
-lemma SUP_extreal_le_addI:
-  assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
-  shows "SUPR UNIV f + y \<le> z"
-proof (cases y)
-  case (real r)
-  then have "\<And>i. f i \<le> z - y" using assms by (simp add: extreal_le_minus_iff)
-  then have "SUPR UNIV f \<le> z - y" by (rule SUP_leI)
-  then show ?thesis using real by (simp add: extreal_le_minus_iff)
-qed (insert assms, auto)
-
-lemma SUPR_extreal_add:
-  fixes f g :: "nat \<Rightarrow> extreal"
-  assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
-  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
-proof (rule extreal_SUPI)
-  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
-  have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
-    unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD)
-  { fix j
-    { fix i
-      have "f i + g j \<le> f i + g (max i j)"
-        using `incseq g`[THEN incseqD] by (rule add_left_mono) auto
-      also have "\<dots> \<le> f (max i j) + g (max i j)"
-        using `incseq f`[THEN incseqD] by (rule add_right_mono) auto
-      also have "\<dots> \<le> y" using * by auto
-      finally have "f i + g j \<le> y" . }
-    then have "SUPR UNIV f + g j \<le> y"
-      using assms(4)[of j] by (intro SUP_extreal_le_addI) auto
-    then have "g j + SUPR UNIV f \<le> y" by (simp add: ac_simps) }
-  then have "SUPR UNIV g + SUPR UNIV f \<le> y"
-    using f by (rule SUP_extreal_le_addI)
-  then show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
-qed (auto intro!: add_mono le_SUPI)
-
-lemma SUPR_extreal_add_pos:
-  fixes f g :: "nat \<Rightarrow> extreal"
-  assumes inc: "incseq f" "incseq g" and pos: "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
-  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
-proof (intro SUPR_extreal_add inc)
-  fix i show "f i \<noteq> -\<infinity>" "g i \<noteq> -\<infinity>" using pos[of i] by auto
-qed
-
-lemma SUPR_extreal_setsum:
-  fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> extreal"
-  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)" and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
-  shows "(SUP i. \<Sum>n\<in>A. f n i) = (\<Sum>n\<in>A. SUPR UNIV (f n))"
-proof cases
-  assume "finite A" then show ?thesis using assms
-    by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_extreal_add_pos)
-qed simp
-
-lemma SUPR_extreal_cmult:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
-  shows "(SUP i. c * f i) = c * SUPR UNIV f"
-proof (rule extreal_SUPI)
-  fix i have "f i \<le> SUPR UNIV f" by (rule le_SUPI) auto
-  then show "c * f i \<le> c * SUPR UNIV f"
-    using `0 \<le> c` by (rule extreal_mult_left_mono)
-next
-  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c * f i \<le> y"
-  show "c * SUPR UNIV f \<le> y"
-  proof cases
-    assume c: "0 < c \<and> c \<noteq> \<infinity>"
-    with * have "SUPR UNIV f \<le> y / c"
-      by (intro SUP_leI) (auto simp: extreal_le_divide_pos)
-    with c show ?thesis
-      by (auto simp: extreal_le_divide_pos)
-  next
-    { assume "c = \<infinity>" have ?thesis
-      proof cases
-        assume "\<forall>i. f i = 0"
-        moreover then have "range f = {0}" by auto
-        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
-      next
-        assume "\<not> (\<forall>i. f i = 0)"
-        then obtain i where "f i \<noteq> 0" by auto
-        with *[of i] `c = \<infinity>` `0 \<le> f i` show ?thesis by (auto split: split_if_asm)
-      qed }
-    moreover assume "\<not> (0 < c \<and> c \<noteq> \<infinity>)"
-    ultimately show ?thesis using * `0 \<le> c` by auto
-  qed
-qed
-
-lemma SUP_PInfty:
-  fixes f :: "'a \<Rightarrow> extreal"
-  assumes "\<And>n::nat. \<exists>i\<in>A. extreal (real n) \<le> f i"
-  shows "(SUP i:A. f i) = \<infinity>"
-  unfolding SUPR_def Sup_eq_top_iff[where 'a=extreal, unfolded top_extreal_def]
-  apply simp
-proof safe
-  fix x assume "x \<noteq> \<infinity>"
-  show "\<exists>i\<in>A. x < f i"
-  proof (cases x)
-    case PInf with `x \<noteq> \<infinity>` show ?thesis by simp
-  next
-    case MInf with assms[of "0"] show ?thesis by force
-  next
-    case (real r)
-    with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < extreal (real n)" by auto
-    moreover from assms[of n] guess i ..
-    ultimately show ?thesis
-      by (auto intro!: bexI[of _ i])
-  qed
-qed
-
-lemma Sup_countable_SUPR:
-  assumes "A \<noteq> {}"
-  shows "\<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
-proof (cases "Sup A")
-  case (real r)
-  have "\<forall>n::nat. \<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
-  proof
-    fix n ::nat have "\<exists>x\<in>A. Sup A - 1 / extreal (real n) < x"
-      using assms real by (intro Sup_extreal_close) (auto simp: one_extreal_def)
-    then guess x ..
-    then show "\<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
-      by (auto intro!: exI[of _ x] simp: extreal_minus_less_iff)
-  qed
-  from choice[OF this] guess f .. note f = this
-  have "SUPR UNIV f = Sup A"
-  proof (rule extreal_SUPI)
-    fix i show "f i \<le> Sup A" using f
-      by (auto intro!: complete_lattice_class.Sup_upper)
-  next
-    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
-    show "Sup A \<le> y"
-    proof (rule extreal_le_epsilon, intro allI impI)
-      fix e :: extreal assume "0 < e"
-      show "Sup A \<le> y + e"
-      proof (cases e)
-        case (real r)
-        hence "0 < r" using `0 < e` by auto
-        then obtain n ::nat where *: "1 / real n < r" "0 < n"
-          using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
-        have "Sup A \<le> f n + 1 / extreal (real n)" using f[THEN spec, of n] by auto
-        also have "1 / extreal (real n) \<le> e" using real * by (auto simp: one_extreal_def )
-        with bound have "f n + 1 / extreal (real n) \<le> y + e" by (rule add_mono) simp
-        finally show "Sup A \<le> y + e" .
-      qed (insert `0 < e`, auto)
-    qed
-  qed
-  with f show ?thesis by (auto intro!: exI[of _ f])
-next
-  case PInf
-  from `A \<noteq> {}` obtain x where "x \<in> A" by auto
-  show ?thesis
-  proof cases
-    assume "\<infinity> \<in> A"
-    moreover then have "\<infinity> \<le> Sup A" by (intro complete_lattice_class.Sup_upper)
-    ultimately show ?thesis by (auto intro!: exI[of _ "\<lambda>x. \<infinity>"])
-  next
-    assume "\<infinity> \<notin> A"
-    have "\<exists>x\<in>A. 0 \<le> x"
-      by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least extreal_infty_less_eq2 linorder_linear)
-    then obtain x where "x \<in> A" "0 \<le> x" by auto
-    have "\<forall>n::nat. \<exists>f. f \<in> A \<and> x + extreal (real n) \<le> f"
-    proof (rule ccontr)
-      assume "\<not> ?thesis"
-      then have "\<exists>n::nat. Sup A \<le> x + extreal (real n)"
-        by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le)
-      then show False using `x \<in> A` `\<infinity> \<notin> A` PInf
-        by(cases x) auto
-    qed
-    from choice[OF this] guess f .. note f = this
-    have "SUPR UNIV f = \<infinity>"
-    proof (rule SUP_PInfty)
-      fix n :: nat show "\<exists>i\<in>UNIV. extreal (real n) \<le> f i"
-        using f[THEN spec, of n] `0 \<le> x`
-        by (cases rule: extreal2_cases[of "f n" x]) (auto intro!: exI[of _ n])
-    qed
-    then show ?thesis using f PInf by (auto intro!: exI[of _ f])
-  qed
-next
-  case MInf
-  with `A \<noteq> {}` have "A = {-\<infinity>}" by (auto simp: Sup_eq_MInfty)
-  then show ?thesis using MInf by (auto intro!: exI[of _ "\<lambda>x. -\<infinity>"])
-qed
-
-lemma SUPR_countable_SUPR:
-  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
-  using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def)
-
-
-lemma Sup_extreal_cadd:
-  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
-  shows "Sup ((\<lambda>x. a + x) ` A) = a + Sup A"
-proof (rule antisym)
-  have *: "\<And>a::extreal. \<And>A. Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A"
-    by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
-  then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
-  show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
-  proof (cases a)
-    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
-  next
-    case (real r)
-    then have **: "op + (- a) ` op + a ` A = A"
-      by (auto simp: image_iff ac_simps zero_extreal_def[symmetric])
-    from *[of "-a" "(\<lambda>x. a + x) ` A"] real show ?thesis unfolding **
-      by (cases rule: extreal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto
-  qed (insert `a \<noteq> -\<infinity>`, auto)
-qed
-
-lemma Sup_extreal_cminus:
-  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
-  shows "Sup ((\<lambda>x. a - x) ` A) = a - Inf A"
-  using Sup_extreal_cadd[of "uminus ` A" a] assms
-  by (simp add: comp_def image_image minus_extreal_def
-                 extreal_Sup_uminus_image_eq)
-
-lemma SUPR_extreal_cminus:
-  fixes A assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
-  shows "(SUP x:A. a - f x) = a - (INF x:A. f x)"
-  using Sup_extreal_cminus[of "f`A" a] assms
-  unfolding SUPR_def INFI_def image_image by auto
-
-lemma Inf_extreal_cminus:
-  fixes A :: "extreal set" assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
-  shows "Inf ((\<lambda>x. a - x) ` A) = a - Sup A"
-proof -
-  { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto }
-  moreover then have "(\<lambda>x. -a - x)`uminus`A = uminus ` (\<lambda>x. a - x) ` A"
-    by (auto simp: image_image)
-  ultimately show ?thesis
-    using Sup_extreal_cminus[of "uminus ` A" "-a"] assms
-    by (auto simp add: extreal_Sup_uminus_image_eq extreal_Inf_uminus_image_eq)
-qed
-
-lemma INFI_extreal_cminus:
-  fixes A assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
-  shows "(INF x:A. a - f x) = a - (SUP x:A. f x)"
-  using Inf_extreal_cminus[of "f`A" a] assms
-  unfolding SUPR_def INFI_def image_image
-  by auto
-
-lemma uminus_extreal_add_uminus_uminus:
-  fixes a b :: extreal shows "a \<noteq> \<infinity> \<Longrightarrow> b \<noteq> \<infinity> \<Longrightarrow> - (- a + - b) = a + b"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma INFI_extreal_add:
-  assumes "decseq f" "decseq g" and fin: "\<And>i. f i \<noteq> \<infinity>" "\<And>i. g i \<noteq> \<infinity>"
-  shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g"
-proof -
-  have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
-    using assms unfolding INF_less_iff by auto
-  { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i"
-      by (rule uminus_extreal_add_uminus_uminus) }
-  then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
-    by simp
-  also have "\<dots> = INFI UNIV f + INFI UNIV g"
-    unfolding extreal_INFI_uminus
-    using assms INF_less
-    by (subst SUPR_extreal_add)
-       (auto simp: extreal_SUPR_uminus intro!: uminus_extreal_add_uminus_uminus)
-  finally show ?thesis .
-qed
-
-subsection "Limits on @{typ extreal}"
-
-subsubsection "Topological space"
-
-instantiation extreal :: topological_space
-begin
-
-definition "open A \<longleftrightarrow> open (extreal -` A)
-       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {extreal x <..} \<subseteq> A))
-       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A))"
-
-lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {extreal x<..} \<subseteq> A)"
-  unfolding open_extreal_def by auto
-
-lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A)"
-  unfolding open_extreal_def by auto
-
-lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{extreal x<..} \<subseteq> A"
-  using open_PInfty[OF assms] by auto
-
-lemma open_MInfty2: assumes "open A" "-\<infinity> \<in> A" obtains x where "{..<extreal x} \<subseteq> A"
-  using open_MInfty[OF assms] by auto
-
-lemma extreal_openE: assumes "open A" obtains x y where
-  "open (extreal -` A)"
-  "\<infinity> \<in> A \<Longrightarrow> {extreal x<..} \<subseteq> A"
-  "-\<infinity> \<in> A \<Longrightarrow> {..<extreal y} \<subseteq> A"
-  using assms open_extreal_def by auto
-
-instance
-proof
-  let ?U = "UNIV::extreal set"
-  show "open ?U" unfolding open_extreal_def
-    by (auto intro!: exI[of _ 0])
-next
-  fix S T::"extreal set" assume "open S" and "open T"
-  from `open S`[THEN extreal_openE] guess xS yS .
-  moreover from `open T`[THEN extreal_openE] guess xT yT .
-  ultimately have
-    "open (extreal -` (S \<inter> T))"
-    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {extreal (max xS xT) <..} \<subseteq> S \<inter> T"
-    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< extreal (min yS yT)} \<subseteq> S \<inter> T"
-    by auto
-  then show "open (S Int T)" unfolding open_extreal_def by blast
-next
-  fix K :: "extreal set set" assume "\<forall>S\<in>K. open S"
-  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (extreal -` S) \<and>
-    (\<infinity> \<in> S \<longrightarrow> {extreal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< extreal y} \<subseteq> S)"
-    by (auto simp: open_extreal_def)
-  then show "open (Union K)" unfolding open_extreal_def
-  proof (intro conjI impI)
-    show "open (extreal -` \<Union>K)"
-      using *[THEN choice] by (auto simp: vimage_Union)
-  qed ((metis UnionE Union_upper subset_trans *)+)
-qed
-end
-
-lemma open_extreal: "open S \<Longrightarrow> open (extreal ` S)"
-  by (auto simp: inj_vimage_image_eq open_extreal_def)
-
-lemma open_extreal_vimage: "open S \<Longrightarrow> open (extreal -` S)"
-  unfolding open_extreal_def by auto
-
-lemma open_extreal_lessThan[intro, simp]: "open {..< a :: extreal}"
-proof -
-  have "\<And>x. extreal -` {..<extreal x} = {..< x}"
-    "extreal -` {..< \<infinity>} = UNIV" "extreal -` {..< -\<infinity>} = {}" by auto
-  then show ?thesis by (cases a) (auto simp: open_extreal_def)
-qed
-
-lemma open_extreal_greaterThan[intro, simp]:
-  "open {a :: extreal <..}"
-proof -
-  have "\<And>x. extreal -` {extreal x<..} = {x<..}"
-    "extreal -` {\<infinity><..} = {}" "extreal -` {-\<infinity><..} = UNIV" by auto
-  then show ?thesis by (cases a) (auto simp: open_extreal_def)
-qed
-
-lemma extreal_open_greaterThanLessThan[intro, simp]: "open {a::extreal <..< b}"
-  unfolding greaterThanLessThan_def by auto
-
-lemma closed_extreal_atLeast[simp, intro]: "closed {a :: extreal ..}"
-proof -
-  have "- {a ..} = {..< a}" by auto
-  then show "closed {a ..}"
-    unfolding closed_def using open_extreal_lessThan by auto
-qed
-
-lemma closed_extreal_atMost[simp, intro]: "closed {.. b :: extreal}"
-proof -
-  have "- {.. b} = {b <..}" by auto
-  then show "closed {.. b}"
-    unfolding closed_def using open_extreal_greaterThan by auto
-qed
-
-lemma closed_extreal_atLeastAtMost[simp, intro]:
-  shows "closed {a :: extreal .. b}"
-  unfolding atLeastAtMost_def by auto
-
-lemma closed_extreal_singleton:
-  "closed {a :: extreal}"
-by (metis atLeastAtMost_singleton closed_extreal_atLeastAtMost)
-
-lemma extreal_open_cont_interval:
-  assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
-  obtains e where "e>0" "{x-e <..< x+e} \<subseteq> S"
-proof-
-  from `open S` have "open (extreal -` S)" by (rule extreal_openE)
-  then obtain e where "0 < e" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> extreal y \<in> S"
-    using assms unfolding open_dist by force
-  show thesis
-  proof (intro that subsetI)
-    show "0 < extreal e" using `0 < e` by auto
-    fix y assume "y \<in> {x - extreal e<..<x + extreal e}"
-    with assms obtain t where "y = extreal t" "dist t (real x) < e"
-      apply (cases y) by (auto simp: dist_real_def)
-    then show "y \<in> S" using e[of t] by auto
-  qed
-qed
-
-lemma extreal_open_cont_interval2:
-  assumes "open S" "x \<in> S" and x: "\<bar>x\<bar> \<noteq> \<infinity>"
-  obtains a b where "a < x" "x < b" "{a <..< b} \<subseteq> S"
-proof-
-  guess e using extreal_open_cont_interval[OF assms] .
-  with that[of "x-e" "x+e"] extreal_between[OF x, of e]
-  show thesis by auto
-qed
-
-instance extreal :: t2_space
-proof
-  fix x y :: extreal assume "x ~= y"
-  let "?P x (y::extreal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
-
-  { fix x y :: extreal assume "x < y"
-    from extreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
-    have "?P x y"
-      apply (rule exI[of _ "{..<z}"])
-      apply (rule exI[of _ "{z<..}"])
-      using z by auto }
-  note * = this
-
-  from `x ~= y`
-  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
-  proof (cases rule: linorder_cases)
-    assume "x = y" with `x ~= y` show ?thesis by simp
-  next assume "x < y" from *[OF this] show ?thesis by auto
-  next assume "y < x" from *[OF this] show ?thesis by auto
-  qed
-qed
-
-subsubsection {* Convergent sequences *}
-
-lemma lim_extreal[simp]:
-  "((\<lambda>n. extreal (f n)) ---> extreal x) net \<longleftrightarrow> (f ---> x) net" (is "?l = ?r")
-proof (intro iffI topological_tendstoI)
-  fix S assume "?l" "open S" "x \<in> S"
-  then show "eventually (\<lambda>x. f x \<in> S) net"
-    using `?l`[THEN topological_tendstoD, OF open_extreal, OF `open S`]
-    by (simp add: inj_image_mem_iff)
-next
-  fix S assume "?r" "open S" "extreal x \<in> S"
-  show "eventually (\<lambda>x. extreal (f x) \<in> S) net"
-    using `?r`[THEN topological_tendstoD, OF open_extreal_vimage, OF `open S`]
-    using `extreal x \<in> S` by auto
-qed
-
-lemma lim_real_of_extreal[simp]:
-  assumes lim: "(f ---> extreal x) net"
-  shows "((\<lambda>x. real (f x)) ---> x) net"
-proof (intro topological_tendstoI)
-  fix S assume "open S" "x \<in> S"
-  then have S: "open S" "extreal x \<in> extreal ` S"
-    by (simp_all add: inj_image_mem_iff)
-  have "\<forall>x. f x \<in> extreal ` S \<longrightarrow> real (f x) \<in> S" by auto
-  from this lim[THEN topological_tendstoD, OF open_extreal, OF S]
-  show "eventually (\<lambda>x. real (f x) \<in> S) net"
-    by (rule eventually_mono)
-qed
-
-lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= extreal B)" (is "?l = ?r")
-proof assume ?r show ?l apply(rule topological_tendstoI)
-    unfolding eventually_sequentially
-  proof- fix S assume "open S" "\<infinity> : S"
-    from open_PInfty[OF this] guess B .. note B=this
-    from `?r`[rule_format,of "B+1"] guess N .. note N=this
-    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
-    proof safe case goal1
-      have "extreal B < extreal (B + 1)" by auto
-      also have "... <= f n" using goal1 N by auto
-      finally show ?case using B by fastsimp
-    qed
-  qed
-next assume ?l show ?r
-  proof fix B::real have "open {extreal B<..}" "\<infinity> : {extreal B<..}" by auto
-    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
-    guess N .. note N=this
-    show "EX N. ALL n>=N. extreal B <= f n" apply(rule_tac x=N in exI) using N by auto
-  qed
-qed
-
-
-lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= extreal B)" (is "?l = ?r")
-proof assume ?r show ?l apply(rule topological_tendstoI)
-    unfolding eventually_sequentially
-  proof- fix S assume "open S" "(-\<infinity>) : S"
-    from open_MInfty[OF this] guess B .. note B=this
-    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
-    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
-    proof safe case goal1
-      have "extreal (B - 1) >= f n" using goal1 N by auto
-      also have "... < extreal B" by auto
-      finally show ?case using B by fastsimp
-    qed
-  qed
-next assume ?l show ?r
-  proof fix B::real have "open {..<extreal B}" "(-\<infinity>) : {..<extreal B}" by auto
-    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
-    guess N .. note N=this
-    show "EX N. ALL n>=N. extreal B >= f n" apply(rule_tac x=N in exI) using N by auto
-  qed
-qed
-
-
-lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= extreal B" shows "l ~= \<infinity>"
-proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
-  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
-  guess N .. note N=this[rule_format,OF le_refl]
-  hence "extreal ?B <= extreal B" using assms(2)[of N] by(rule order_trans)
-  hence "extreal ?B < extreal ?B" apply (rule le_less_trans) by auto
-  thus False by auto
-qed
-
-
-lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= extreal B" shows "l ~= (-\<infinity>)"
-proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
-  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
-  guess N .. note N=this[rule_format,OF le_refl]
-  hence "extreal B <= extreal ?B" using assms(2)[of N] order_trans[of "extreal B" "f N" "extreal(B - 1)"] by blast
-  thus False by auto
-qed
-
-
-lemma tendsto_explicit:
-  "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
-  unfolding tendsto_def eventually_sequentially by auto
-
-
-lemma tendsto_obtains_N:
-  assumes "f ----> f0"
-  assumes "open S" "f0 : S"
-  obtains N where "ALL n>=N. f n : S"
-  using tendsto_explicit[of f f0] assms by auto
-
-
-lemma tail_same_limit:
-  fixes X Y N
-  assumes "X ----> L" "ALL n>=N. X n = Y n"
-  shows "Y ----> L"
-proof-
-{ fix S assume "open S" and "L:S"
-  from this obtain N1 where "ALL n>=N1. X n : S"
-     using assms unfolding tendsto_def eventually_sequentially by auto
-  hence "ALL n>=max N N1. Y n : S" using assms by auto
-  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
-}
-thus ?thesis using tendsto_explicit by auto
-qed
-
-
-lemma Lim_bounded_PInfty2:
-assumes lim:"f ----> l" and "ALL n>=N. f n <= extreal B"
-shows "l ~= \<infinity>"
-proof-
-  def g == "(%n. if n>=N then f n else extreal B)"
-  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
-  moreover have "!!n. g n <= extreal B" using g_def assms by auto
-  ultimately show ?thesis using  Lim_bounded_PInfty by auto
-qed
-
-lemma Lim_bounded_extreal:
-  assumes lim:"f ----> (l :: extreal)"
-  and "ALL n>=M. f n <= C"
-  shows "l<=C"
-proof-
-{ assume "l=(-\<infinity>)" hence ?thesis by auto }
-moreover
-{ assume "~(l=(-\<infinity>))"
-  { assume "C=\<infinity>" hence ?thesis by auto }
-  moreover
-  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
-    hence "l=(-\<infinity>)" using assms
-       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
-    hence ?thesis by auto }
-  moreover
-  { assume "EX B. C = extreal B"
-    from this obtain B where B_def: "C=extreal B" by auto
-    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
-    from this obtain m where m_def: "extreal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
-    from this obtain N where N_def: "ALL n>=N. f n : {extreal(m - 1) <..< extreal(m+1)}"
-       apply (subst tendsto_obtains_N[of f l "{extreal(m - 1) <..< extreal(m+1)}"]) using assms by auto
-    { fix n assume "n>=N"
-      hence "EX r. extreal r = f n" using N_def by (cases "f n") auto
-    } from this obtain g where g_def: "ALL n>=N. extreal (g n) = f n" by metis
-    hence "(%n. extreal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
-    hence *: "(%n. g n) ----> m" using m_def by auto
-    { fix n assume "n>=max N M"
-      hence "extreal (g n) <= extreal B" using assms g_def B_def by auto
-      hence "g n <= B" by auto
-    } hence "EX N. ALL n>=N. g n <= B" by blast
-    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
-    hence ?thesis using m_def B_def by auto
-  } ultimately have ?thesis by (cases C) auto
-} ultimately show ?thesis by blast
-qed
-
-lemma real_of_extreal_mult[simp]:
-  fixes a b :: extreal shows "real (a * b) = real a * real b"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma real_of_extreal_eq_0:
-  "real x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity> \<or> x = 0"
-  by (cases x) auto
-
-lemma tendsto_extreal_realD:
-  fixes f :: "'a \<Rightarrow> extreal"
-  assumes "x \<noteq> 0" and tendsto: "((\<lambda>x. extreal (real (f x))) ---> x) net"
-  shows "(f ---> x) net"
-proof (intro topological_tendstoI)
-  fix S assume S: "open S" "x \<in> S"
-  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}" by auto
-  from tendsto[THEN topological_tendstoD, OF this]
-  show "eventually (\<lambda>x. f x \<in> S) net"
-    by (rule eventually_rev_mp) (auto simp: extreal_real real_of_extreal_0)
-qed
-
-lemma tendsto_extreal_realI:
-  fixes f :: "'a \<Rightarrow> extreal"
-  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>" and tendsto: "(f ---> x) net"
-  shows "((\<lambda>x. extreal (real (f x))) ---> x) net"
-proof (intro topological_tendstoI)
-  fix S assume "open S" "x \<in> S"
-  with x have "open (S - {\<infinity>, -\<infinity>})" "x \<in> S - {\<infinity>, -\<infinity>}" by auto
-  from tendsto[THEN topological_tendstoD, OF this]
-  show "eventually (\<lambda>x. extreal (real (f x)) \<in> S) net"
-    by (elim eventually_elim1) (auto simp: extreal_real)
-qed
-
-lemma extreal_mult_cancel_left:
-  fixes a b c :: extreal shows "a * b = a * c \<longleftrightarrow>
-    ((\<bar>a\<bar> = \<infinity> \<and> 0 < b * c) \<or> a = 0 \<or> b = c)"
-  by (cases rule: extreal3_cases[of a b c])
-     (simp_all add: zero_less_mult_iff)
-
-lemma extreal_inj_affinity:
-  assumes "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" "\<bar>t\<bar> \<noteq> \<infinity>"
-  shows "inj_on (\<lambda>x. m * x + t) A"
-  using assms
-  by (cases rule: extreal2_cases[of m t])
-     (auto intro!: inj_onI simp: extreal_add_cancel_right extreal_mult_cancel_left)
-
-lemma extreal_PInfty_eq_plus[simp]:
-  shows "\<infinity> = a + b \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_MInfty_eq_plus[simp]:
-  shows "-\<infinity> = a + b \<longleftrightarrow> (a = -\<infinity> \<and> b \<noteq> \<infinity>) \<or> (b = -\<infinity> \<and> a \<noteq> \<infinity>)"
-  by (cases rule: extreal2_cases[of a b]) auto
-
-lemma extreal_less_divide_pos:
-  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y < z / x \<longleftrightarrow> x * y < z"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_divide_less_pos:
-  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y / x < z \<longleftrightarrow> y < x * z"
-  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
-
-lemma extreal_divide_eq:
-  "b \<noteq> 0 \<Longrightarrow> \<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> a / b = c \<longleftrightarrow> a = b * c"
-  by (cases rule: extreal3_cases[of a b c])
-     (simp_all add: field_simps)
-
-lemma extreal_inverse_not_MInfty[simp]: "inverse a \<noteq> -\<infinity>"
-  by (cases a) auto
-
-lemma extreal_mult_m1[simp]: "x * extreal (-1) = -x"
-  by (cases x) auto
-
-lemma extreal_LimI_finite:
-  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
-  assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r"
-  shows "u ----> x"
-proof (rule topological_tendstoI, unfold eventually_sequentially)
-  obtain rx where rx_def: "x=extreal rx" using assms by (cases x) auto
-  fix S assume "open S" "x : S"
-  then have "open (extreal -` S)" unfolding open_extreal_def by auto
-  with `x \<in> S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> extreal y \<in> S"
-    unfolding open_real_def rx_def by auto
-  then obtain n where
-    upper: "!!N. n <= N ==> u N < x + extreal r" and
-    lower: "!!N. n <= N ==> x < u N + extreal r" using assms(2)[of "extreal r"] by auto
-  show "EX N. ALL n>=N. u n : S"
-  proof (safe intro!: exI[of _ n])
-    fix N assume "n <= N"
-    from upper[OF this] lower[OF this] assms `0 < r`
-    have "u N ~: {\<infinity>,(-\<infinity>)}" by auto
-    from this obtain ra where ra_def: "(u N) = extreal ra" by (cases "u N") auto
-    hence "rx < ra + r" and "ra < rx + r"
-       using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto
-    hence "dist (real (u N)) rx < r"
-      using rx_def ra_def
-      by (auto simp: dist_real_def abs_diff_less_iff field_simps)
-    from dist[OF this] show "u N : S" using `u N  ~: {\<infinity>, -\<infinity>}`
-      by (auto simp: extreal_real split: split_if_asm)
-  qed
-qed
-
-lemma extreal_LimI_finite_iff:
-  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
-  shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))"
-  (is "?lhs <-> ?rhs")
-proof
-  assume lim: "u ----> x"
-  { fix r assume "(r::extreal)>0"
-    from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}"
-       apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
-       using lim extreal_between[of x r] assms `r>0` by auto
-    hence "EX N. ALL n>=N. u n < x + r & x < u n + r"
-      using extreal_minus_less[of r x] by (cases r) auto
-  } then show "?rhs" by auto
-next
-  assume ?rhs then show "u ----> x"
-    using extreal_LimI_finite[of x] assms by auto
-qed
-
-
-subsubsection {* @{text Liminf} and @{text Limsup} *}
-
-definition
-  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
-
-definition
-  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
-
-lemma Liminf_Sup:
-  fixes f :: "'a => 'b::{complete_lattice, linorder}"
-  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
-  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
-
-lemma Limsup_Inf:
-  fixes f :: "'a => 'b::{complete_lattice, linorder}"
-  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
-  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
-
-lemma extreal_SupI:
-  fixes x :: extreal
-  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
-  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
-  shows "Sup A = x"
-  unfolding Sup_extreal_def
-  using assms by (auto intro!: Least_equality)
-
-lemma extreal_InfI:
-  fixes x :: extreal
-  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
-  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
-  shows "Inf A = x"
-  unfolding Inf_extreal_def
-  using assms by (auto intro!: Greatest_equality)
-
-lemma Limsup_const:
-  fixes c :: "'a::{complete_lattice, linorder}"
-  assumes ntriv: "\<not> trivial_limit net"
-  shows "Limsup net (\<lambda>x. c) = c"
-  unfolding Limsup_Inf
-proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
-  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
-  show "c \<le> x"
-  proof (rule ccontr)
-    assume "\<not> c \<le> x" then have "x < c" by auto
-    then show False using ntriv * by (auto simp: trivial_limit_def)
-  qed
-qed auto
-
-lemma Liminf_const:
-  fixes c :: "'a::{complete_lattice, linorder}"
-  assumes ntriv: "\<not> trivial_limit net"
-  shows "Liminf net (\<lambda>x. c) = c"
-  unfolding Liminf_Sup
-proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
-  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
-  show "x \<le> c"
-  proof (rule ccontr)
-    assume "\<not> x \<le> c" then have "c < x" by auto
-    then show False using ntriv * by (auto simp: trivial_limit_def)
-  qed
-qed auto
-
-lemma mono_set:
-  fixes S :: "('a::order) set"
-  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
-  by (auto simp: mono_def mem_def)
-
-lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
-lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
-lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
-lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
-
-lemma mono_set_iff:
-  fixes S :: "'a::{linorder,complete_lattice} set"
-  defines "a \<equiv> Inf S"
-  shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
-proof
-  assume "mono S"
-  then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set)
-  show ?c
-  proof cases
-    assume "a \<in> S"
-    show ?c
-      using mono[OF _ `a \<in> S`]
-      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
-  next
-    assume "a \<notin> S"
-    have "S = {a <..}"
-    proof safe
-      fix x assume "x \<in> S"
-      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
-      then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
-    next
-      fix x assume "a < x"
-      then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff ..
-      with mono[of y x] show "x \<in> S" by auto
-    qed
-    then show ?c ..
-  qed
-qed auto
-
-lemma lim_imp_Liminf:
-  fixes f :: "'a \<Rightarrow> extreal"
-  assumes ntriv: "\<not> trivial_limit net"
-  assumes lim: "(f ---> f0) net"
-  shows "Liminf net f = f0"
-  unfolding Liminf_Sup
-proof (safe intro!: extreal_SupI)
-  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
-  show "y \<le> f0"
-  proof (rule extreal_le_extreal)
-    fix B assume "B < y"
-    { assume "f0 < B"
-      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
-         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
-         by (auto intro: eventually_conj)
-      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
-      finally have False using ntriv[unfolded trivial_limit_def] by auto
-    } then show "B \<le> f0" by (metis linorder_le_less_linear)
-  qed
-next
-  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
-  show "f0 \<le> y"
-  proof (safe intro!: *[rule_format])
-    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
-      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
-  qed
-qed
-
-lemma extreal_Liminf_le_Limsup:
-  fixes f :: "'a \<Rightarrow> extreal"
-  assumes ntriv: "\<not> trivial_limit net"
-  shows "Liminf net f \<le> Limsup net f"
-  unfolding Limsup_Inf Liminf_Sup
-proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
-  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
-  show "u \<le> v"
-  proof (rule ccontr)
-    assume "\<not> u \<le> v"
-    then obtain t where "t < u" "v < t"
-      using extreal_dense[of v u] by (auto simp: not_le)
-    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
-      using * by (auto intro: eventually_conj)
-    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
-    finally show False using ntriv by (auto simp: trivial_limit_def)
-  qed
-qed
-
-lemma Liminf_mono:
-  fixes f g :: "'a => extreal"
-  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
-  shows "Liminf net f \<le> Liminf net g"
-  unfolding Liminf_Sup
-proof (safe intro!: Sup_mono bexI)
-  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
-  then have "eventually (\<lambda>x. y < f x) net" by auto
-  then show "eventually (\<lambda>x. y < g x) net"
-    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
-qed simp
-
-lemma Liminf_eq:
-  fixes f g :: "'a \<Rightarrow> extreal"
-  assumes "eventually (\<lambda>x. f x = g x) net"
-  shows "Liminf net f = Liminf net g"
-  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
-
-lemma Liminf_mono_all:
-  fixes f g :: "'a \<Rightarrow> extreal"
-  assumes "\<And>x. f x \<le> g x"
-  shows "Liminf net f \<le> Liminf net g"
-  using assms by (intro Liminf_mono always_eventually) auto
-
-lemma Limsup_mono:
-  fixes f g :: "'a \<Rightarrow> extreal"
-  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
-  shows "Limsup net f \<le> Limsup net g"
-  unfolding Limsup_Inf
-proof (safe intro!: Inf_mono bexI)
-  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
-  then have "eventually (\<lambda>x. g x < y) net" by auto
-  then show "eventually (\<lambda>x. f x < y) net"
-    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
-qed simp
-
-lemma Limsup_mono_all:
-  fixes f g :: "'a \<Rightarrow> extreal"
-  assumes "\<And>x. f x \<le> g x"
-  shows "Limsup net f \<le> Limsup net g"
-  using assms by (intro Limsup_mono always_eventually) auto
-
-lemma Limsup_eq:
-  fixes f g :: "'a \<Rightarrow> extreal"
-  assumes "eventually (\<lambda>x. f x = g x) net"
-  shows "Limsup net f = Limsup net g"
-  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
-
-abbreviation "liminf \<equiv> Liminf sequentially"
-
-abbreviation "limsup \<equiv> Limsup sequentially"
-
-lemma (in complete_lattice) less_INFD:
-  assumes "y < INFI A f"" i \<in> A" shows "y < f i"
-proof -
-  note `y < INFI A f`
-  also have "INFI A f \<le> f i" using `i \<in> A` by (rule INF_leI)
-  finally show "y < f i" .
-qed
-
-lemma liminf_SUPR_INFI:
-  fixes f :: "nat \<Rightarrow> extreal"
-  shows "liminf f = (SUP n. INF m:{n..}. f m)"
-  unfolding Liminf_Sup eventually_sequentially
-proof (safe intro!: antisym complete_lattice_class.Sup_least)
-  fix x assume *: "\<forall>y<x. \<exists>N. \<forall>n\<ge>N. y < f n" show "x \<le> (SUP n. INF m:{n..}. f m)"
-  proof (rule extreal_le_extreal)
-    fix y assume "y < x"
-    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
-    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
-    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto
-    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
-  qed
-next
-  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
-  proof (unfold SUPR_def, safe intro!: Sup_mono bexI)
-    fix y n assume "y < INFI {n..} f"
-    from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
-  qed (rule order_refl)
-qed
-
-lemma tail_same_limsup:
-  fixes X Y :: "nat => extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
-  shows "limsup X = limsup Y"
-  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
-
-lemma tail_same_liminf:
-  fixes X Y :: "nat => extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
-  shows "liminf X = liminf Y"
-  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
-
-lemma liminf_mono:
-  fixes X Y :: "nat \<Rightarrow> extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
-  shows "liminf X \<le> liminf Y"
-  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
-
-lemma limsup_mono:
-  fixes X Y :: "nat => extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
-  shows "limsup X \<le> limsup Y"
-  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
-
-declare trivial_limit_sequentially[simp]
-
-lemma
-  fixes X :: "nat \<Rightarrow> extreal"
-  shows extreal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
-    and extreal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
-  unfolding incseq_def decseq_def by auto
-
-lemma liminf_bounded:
-  fixes X Y :: "nat \<Rightarrow> extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
-  shows "C \<le> liminf X"
-  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
-
-lemma limsup_bounded:
-  fixes X Y :: "nat => extreal"
-  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
-  shows "limsup X \<le> C"
-  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
-
-lemma liminf_bounded_iff:
-  fixes x :: "nat \<Rightarrow> extreal"
-  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
-proof safe
-  fix B assume "B < C" "C \<le> liminf x"
-  then have "B < liminf x" by auto
-  then obtain N where "B < (INF m:{N..}. x m)"
-    unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto
-  from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
-next
-  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
-  { fix B assume "B<C"
-    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
-    hence "B \<le> (INF m:{N..}. x m)" by (intro le_INFI) auto
-    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp
-    finally have "B \<le> liminf x" .
-  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
-qed
-
-lemma liminf_subseq_mono:
-  fixes X :: "nat \<Rightarrow> extreal"
-  assumes "subseq r"
-  shows "liminf X \<le> liminf (X \<circ> r) "
-proof-
-  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
-  proof (safe intro!: INF_mono)
-    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
-      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
-  qed
-  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
-qed
-
-lemma extreal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "extreal (real x) = x"
-  using assms by auto
-
-lemma extreal_le_extreal_bounded:
-  fixes x y z :: extreal
-  assumes "z \<le> y"
-  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
-  shows "x \<le> y"
-proof (rule extreal_le_extreal)
-  fix B assume "B < x"
-  show "B \<le> y"
-  proof cases
-    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
-  next
-    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
-  qed
-qed
-
-lemma fixes x y :: extreal
-  shows Sup_atMost[simp]: "Sup {.. y} = y"
-    and Sup_lessThan[simp]: "Sup {..< y} = y"
-    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
-    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
-    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
-  by (auto simp: Sup_extreal_def intro!: Least_equality
-           intro: extreal_le_extreal extreal_le_extreal_bounded[of x])
-
-lemma Sup_greaterThanlessThan[simp]:
-  fixes x y :: extreal assumes "x < y" shows "Sup { x <..< y} = y"
-  unfolding Sup_extreal_def
-proof (intro Least_equality extreal_le_extreal_bounded[of _ _ y])
-  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
-  from extreal_dense[OF `x < y`] guess w .. note w = this
-  with z[THEN bspec, of w] show "x \<le> z" by auto
-qed auto
-
-lemma real_extreal_id: "real o extreal = id"
-proof-
-{ fix x have "(real o extreal) x = id x" by auto }
-from this show ?thesis using ext by blast
-qed
-
-lemma open_image_extreal: "open(UNIV-{\<infinity>,(-\<infinity>)})"
-by (metis range_extreal open_extreal open_UNIV)
-
-lemma extreal_le_distrib:
-  fixes a b c :: extreal shows "c * (a + b) \<le> c * a + c * b"
-  by (cases rule: extreal3_cases[of a b c])
-     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
-
-lemma extreal_pos_distrib:
-  fixes a b c :: extreal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
-  using assms by (cases rule: extreal3_cases[of a b c])
-                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
-
-lemma extreal_pos_le_distrib:
-fixes a b c :: extreal
-assumes "c>=0"
-shows "c * (a + b) <= c * a + c * b"
-  using assms by (cases rule: extreal3_cases[of a b c])
-                 (auto simp add: field_simps)
-
-lemma extreal_max_mono:
-  "[| (a::extreal) <= b; c <= d |] ==> max a c <= max b d"
-  by (metis sup_extreal_def sup_mono)
-
-
-lemma extreal_max_least:
-  "[| (a::extreal) <= x; c <= x |] ==> max a c <= x"
-  by (metis sup_extreal_def sup_least)
-
-end
--- a/src/HOL/Library/Library.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Library/Library.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -15,6 +15,7 @@
   Diagonalize
   Dlist_Cset
   Eval_Witness
+  Extended_Nat
   Float
   Formal_Power_Series
   Fraction_Field
@@ -35,7 +36,6 @@
   Monad_Syntax
   More_List
   Multiset
-  Nat_Infinity
   Nested_Environment
   Numeral_Type
   OptionalSugar
--- a/src/HOL/Library/Nat_Infinity.thy	Wed Jul 20 10:11:08 2011 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,551 +0,0 @@
-(*  Title:      HOL/Library/Nat_Infinity.thy
-    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
-    Contributions: David Trachtenherz, TU Muenchen
-*)
-
-header {* Natural numbers with infinity *}
-
-theory Nat_Infinity
-imports Main
-begin
-
-subsection {* Type definition *}
-
-text {*
-  We extend the standard natural numbers by a special value indicating
-  infinity.
-*}
-
-datatype inat = Fin nat | Infty
-
-notation (xsymbols)
-  Infty  ("\<infinity>")
-
-notation (HTML output)
-  Infty  ("\<infinity>")
-
-
-lemma not_Infty_eq[iff]: "(x ~= Infty) = (EX i. x = Fin i)"
-by (cases x) auto
-
-lemma not_Fin_eq [iff]: "(ALL y. x ~= Fin y) = (x = Infty)"
-by (cases x) auto
-
-
-primrec the_Fin :: "inat \<Rightarrow> nat"
-where "the_Fin (Fin n) = n"
-
-
-subsection {* Constructors and numbers *}
-
-instantiation inat :: "{zero, one, number}"
-begin
-
-definition
-  "0 = Fin 0"
-
-definition
-  [code_unfold]: "1 = Fin 1"
-
-definition
-  [code_unfold, code del]: "number_of k = Fin (number_of k)"
-
-instance ..
-
-end
-
-definition iSuc :: "inat \<Rightarrow> inat" where
-  "iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
-
-lemma Fin_0: "Fin 0 = 0"
-  by (simp add: zero_inat_def)
-
-lemma Fin_1: "Fin 1 = 1"
-  by (simp add: one_inat_def)
-
-lemma Fin_number: "Fin (number_of k) = number_of k"
-  by (simp add: number_of_inat_def)
-
-lemma one_iSuc: "1 = iSuc 0"
-  by (simp add: zero_inat_def one_inat_def iSuc_def)
-
-lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
-  by (simp add: zero_inat_def)
-
-lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
-  by (simp add: zero_inat_def)
-
-lemma zero_inat_eq [simp]:
-  "number_of k = (0\<Colon>inat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
-  "(0\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
-  unfolding zero_inat_def number_of_inat_def by simp_all
-
-lemma one_inat_eq [simp]:
-  "number_of k = (1\<Colon>inat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
-  "(1\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
-  unfolding one_inat_def number_of_inat_def by simp_all
-
-lemma zero_one_inat_neq [simp]:
-  "\<not> 0 = (1\<Colon>inat)"
-  "\<not> 1 = (0\<Colon>inat)"
-  unfolding zero_inat_def one_inat_def by simp_all
-
-lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1"
-  by (simp add: one_inat_def)
-
-lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>"
-  by (simp add: one_inat_def)
-
-lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k"
-  by (simp add: number_of_inat_def)
-
-lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>"
-  by (simp add: number_of_inat_def)
-
-lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)"
-  by (simp add: iSuc_def)
-
-lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))"
-  by (simp add: iSuc_Fin number_of_inat_def)
-
-lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
-  by (simp add: iSuc_def)
-
-lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
-  by (simp add: iSuc_def zero_inat_def split: inat.splits)
-
-lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
-  by (rule iSuc_ne_0 [symmetric])
-
-lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
-  by (simp add: iSuc_def split: inat.splits)
-
-lemma number_of_inat_inject [simp]:
-  "(number_of k \<Colon> inat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
-  by (simp add: number_of_inat_def)
-
-
-subsection {* Addition *}
-
-instantiation inat :: comm_monoid_add
-begin
-
-definition [nitpick_simp]:
-  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))"
-
-lemma plus_inat_simps [simp, code]:
-  "Fin m + Fin n = Fin (m + n)"
-  "\<infinity> + q = \<infinity>"
-  "q + \<infinity> = \<infinity>"
-  by (simp_all add: plus_inat_def split: inat.splits)
-
-instance proof
-  fix n m q :: inat
-  show "n + m + q = n + (m + q)"
-    by (cases n, auto, cases m, auto, cases q, auto)
-  show "n + m = m + n"
-    by (cases n, auto, cases m, auto)
-  show "0 + n = n"
-    by (cases n) (simp_all add: zero_inat_def)
-qed
-
-end
-
-lemma plus_inat_0 [simp]:
-  "0 + (q\<Colon>inat) = q"
-  "(q\<Colon>inat) + 0 = q"
-  by (simp_all add: plus_inat_def zero_inat_def split: inat.splits)
-
-lemma plus_inat_number [simp]:
-  "(number_of k \<Colon> inat) + number_of l = (if k < Int.Pls then number_of l
-    else if l < Int.Pls then number_of k else number_of (k + l))"
-  unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
-
-lemma iSuc_number [simp]:
-  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
-  unfolding iSuc_number_of
-  unfolding one_inat_def number_of_inat_def Suc_nat_number_of if_distrib [symmetric] ..
-
-lemma iSuc_plus_1:
-  "iSuc n = n + 1"
-  by (cases n) (simp_all add: iSuc_Fin one_inat_def)
-  
-lemma plus_1_iSuc:
-  "1 + q = iSuc q"
-  "q + 1 = iSuc q"
-by (simp_all add: iSuc_plus_1 add_ac)
-
-lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
-by (simp_all add: iSuc_plus_1 add_ac)
-
-lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
-by (simp only: add_commute[of m] iadd_Suc)
-
-lemma iadd_is_0: "(m + n = (0::inat)) = (m = 0 \<and> n = 0)"
-by (cases m, cases n, simp_all add: zero_inat_def)
-
-subsection {* Multiplication *}
-
-instantiation inat :: comm_semiring_1
-begin
-
-definition times_inat_def [nitpick_simp]:
-  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
-    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
-
-lemma times_inat_simps [simp, code]:
-  "Fin m * Fin n = Fin (m * n)"
-  "\<infinity> * \<infinity> = \<infinity>"
-  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
-  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
-  unfolding times_inat_def zero_inat_def
-  by (simp_all split: inat.split)
-
-instance proof
-  fix a b c :: inat
-  show "(a * b) * c = a * (b * c)"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "a * b = b * a"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "1 * a = a"
-    unfolding times_inat_def zero_inat_def one_inat_def
-    by (simp split: inat.split)
-  show "(a + b) * c = a * c + b * c"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split add: left_distrib)
-  show "0 * a = 0"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "a * 0 = 0"
-    unfolding times_inat_def zero_inat_def
-    by (simp split: inat.split)
-  show "(0::inat) \<noteq> 1"
-    unfolding zero_inat_def one_inat_def
-    by simp
-qed
-
-end
-
-lemma mult_iSuc: "iSuc m * n = n + m * n"
-  unfolding iSuc_plus_1 by (simp add: algebra_simps)
-
-lemma mult_iSuc_right: "m * iSuc n = m + m * n"
-  unfolding iSuc_plus_1 by (simp add: algebra_simps)
-
-lemma of_nat_eq_Fin: "of_nat n = Fin n"
-  apply (induct n)
-  apply (simp add: Fin_0)
-  apply (simp add: plus_1_iSuc iSuc_Fin)
-  done
-
-instance inat :: number_semiring
-proof
-  fix n show "number_of (int n) = (of_nat n :: inat)"
-    unfolding number_of_inat_def number_of_int of_nat_id of_nat_eq_Fin ..
-qed
-
-instance inat :: semiring_char_0 proof
-  have "inj Fin" by (rule injI) simp
-  then show "inj (\<lambda>n. of_nat n :: inat)" by (simp add: of_nat_eq_Fin)
-qed
-
-lemma imult_is_0[simp]: "((m::inat) * n = 0) = (m = 0 \<or> n = 0)"
-by(auto simp add: times_inat_def zero_inat_def split: inat.split)
-
-lemma imult_is_Infty: "((a::inat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
-by(auto simp add: times_inat_def zero_inat_def split: inat.split)
-
-
-subsection {* Subtraction *}
-
-instantiation inat :: minus
-begin
-
-definition diff_inat_def:
-"a - b = (case a of (Fin x) \<Rightarrow> (case b of (Fin y) \<Rightarrow> Fin (x - y) | \<infinity> \<Rightarrow> 0)
-          | \<infinity> \<Rightarrow> \<infinity>)"
-
-instance ..
-
-end
-
-lemma idiff_Fin_Fin[simp,code]: "Fin a - Fin b = Fin (a - b)"
-by(simp add: diff_inat_def)
-
-lemma idiff_Infty[simp,code]: "\<infinity> - n = \<infinity>"
-by(simp add: diff_inat_def)
-
-lemma idiff_Infty_right[simp,code]: "Fin a - \<infinity> = 0"
-by(simp add: diff_inat_def)
-
-lemma idiff_0[simp]: "(0::inat) - n = 0"
-by (cases n, simp_all add: zero_inat_def)
-
-lemmas idiff_Fin_0[simp] = idiff_0[unfolded zero_inat_def]
-
-lemma idiff_0_right[simp]: "(n::inat) - 0 = n"
-by (cases n) (simp_all add: zero_inat_def)
-
-lemmas idiff_Fin_0_right[simp] = idiff_0_right[unfolded zero_inat_def]
-
-lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::inat) - n = 0"
-by(auto simp: zero_inat_def)
-
-lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
-by(simp add: iSuc_def split: inat.split)
-
-lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
-by(simp add: one_inat_def iSuc_Fin[symmetric] zero_inat_def[symmetric])
-
-(*lemmas idiff_self_eq_0_Fin = idiff_self_eq_0[unfolded zero_inat_def]*)
-
-
-subsection {* Ordering *}
-
-instantiation inat :: linordered_ab_semigroup_add
-begin
-
-definition [nitpick_simp]:
-  "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
-    | \<infinity> \<Rightarrow> True)"
-
-definition [nitpick_simp]:
-  "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
-    | \<infinity> \<Rightarrow> False)"
-
-lemma inat_ord_simps [simp]:
-  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
-  "Fin m < Fin n \<longleftrightarrow> m < n"
-  "q \<le> \<infinity>"
-  "q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>"
-  "\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>"
-  "\<infinity> < q \<longleftrightarrow> False"
-  by (simp_all add: less_eq_inat_def less_inat_def split: inat.splits)
-
-lemma inat_ord_code [code]:
-  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
-  "Fin m < Fin n \<longleftrightarrow> m < n"
-  "q \<le> \<infinity> \<longleftrightarrow> True"
-  "Fin m < \<infinity> \<longleftrightarrow> True"
-  "\<infinity> \<le> Fin n \<longleftrightarrow> False"
-  "\<infinity> < q \<longleftrightarrow> False"
-  by simp_all
-
-instance by default
-  (auto simp add: less_eq_inat_def less_inat_def plus_inat_def split: inat.splits)
-
-end
-
-instance inat :: ordered_comm_semiring
-proof
-  fix a b c :: inat
-  assume "a \<le> b" and "0 \<le> c"
-  thus "c * a \<le> c * b"
-    unfolding times_inat_def less_eq_inat_def zero_inat_def
-    by (simp split: inat.splits)
-qed
-
-lemma inat_ord_number [simp]:
-  "(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
-  "(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
-  by (simp_all add: number_of_inat_def)
-
-lemma i0_lb [simp]: "(0\<Colon>inat) \<le> n"
-  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma ile0_eq [simp]: "n \<le> (0\<Colon>inat) \<longleftrightarrow> n = 0"
-by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R"
-  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
-
-lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R"
-  by simp
-
-lemma not_iless0 [simp]: "\<not> n < (0\<Colon>inat)"
-  by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma i0_less [simp]: "(0\<Colon>inat) < n \<longleftrightarrow> n \<noteq> 0"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
-  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
- 
-lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
-  by (simp add: iSuc_def less_inat_def split: inat.splits)
-
-lemma ile_iSuc [simp]: "n \<le> iSuc n"
-  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
-
-lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
-  by (simp add: zero_inat_def iSuc_def less_eq_inat_def split: inat.splits)
-
-lemma i0_iless_iSuc [simp]: "0 < iSuc n"
-  by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.splits)
-
-lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
-by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.split)
-
-lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
-  by (simp add: iSuc_def less_eq_inat_def less_inat_def split: inat.splits)
-
-lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n"
-  by (cases n) auto
-
-lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n"
-  by (auto simp add: iSuc_def less_inat_def split: inat.splits)
-
-lemma imult_Infty: "(0::inat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma imult_Infty_right: "(0::inat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
-by (simp add: zero_inat_def less_inat_def split: inat.splits)
-
-lemma inat_0_less_mult_iff: "(0 < (m::inat) * n) = (0 < m \<and> 0 < n)"
-by (simp only: i0_less imult_is_0, simp)
-
-lemma mono_iSuc: "mono iSuc"
-by(simp add: mono_def)
-
-
-lemma min_inat_simps [simp]:
-  "min (Fin m) (Fin n) = Fin (min m n)"
-  "min q 0 = 0"
-  "min 0 q = 0"
-  "min q \<infinity> = q"
-  "min \<infinity> q = q"
-  by (auto simp add: min_def)
-
-lemma max_inat_simps [simp]:
-  "max (Fin m) (Fin n) = Fin (max m n)"
-  "max q 0 = q"
-  "max 0 q = q"
-  "max q \<infinity> = \<infinity>"
-  "max \<infinity> q = \<infinity>"
-  by (simp_all add: max_def)
-
-lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k"
-  by (cases n) simp_all
-
-lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k"
-  by (cases n) simp_all
-
-lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
-apply (induct_tac k)
- apply (simp (no_asm) only: Fin_0)
- apply (fast intro: le_less_trans [OF i0_lb])
-apply (erule exE)
-apply (drule spec)
-apply (erule exE)
-apply (drule ileI1)
-apply (rule iSuc_Fin [THEN subst])
-apply (rule exI)
-apply (erule (1) le_less_trans)
-done
-
-instantiation inat :: "{bot, top}"
-begin
-
-definition bot_inat :: inat where
-  "bot_inat = 0"
-
-definition top_inat :: inat where
-  "top_inat = \<infinity>"
-
-instance proof
-qed (simp_all add: bot_inat_def top_inat_def)
-
-end
-
-lemma finite_Fin_bounded:
-  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> Fin n"
-  shows "finite A"
-proof (rule finite_subset)
-  show "finite (Fin ` {..n})" by blast
-
-  have "A \<subseteq> {..Fin n}" using le_fin by fastsimp
-  also have "\<dots> \<subseteq> Fin ` {..n}"
-    by (rule subsetI) (case_tac x, auto)
-  finally show "A \<subseteq> Fin ` {..n}" .
-qed
-
-
-subsection {* Well-ordering *}
-
-lemma less_FinE:
-  "[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P"
-by (induct n) auto
-
-lemma less_InftyE:
-  "[| n < Infty; !!k. n = Fin k ==> P |] ==> P"
-by (induct n) auto
-
-lemma inat_less_induct:
-  assumes prem: "!!n. \<forall>m::inat. m < n --> P m ==> P n" shows "P n"
-proof -
-  have P_Fin: "!!k. P (Fin k)"
-    apply (rule nat_less_induct)
-    apply (rule prem, clarify)
-    apply (erule less_FinE, simp)
-    done
-  show ?thesis
-  proof (induct n)
-    fix nat
-    show "P (Fin nat)" by (rule P_Fin)
-  next
-    show "P Infty"
-      apply (rule prem, clarify)
-      apply (erule less_InftyE)
-      apply (simp add: P_Fin)
-      done
-  qed
-qed
-
-instance inat :: wellorder
-proof
-  fix P and n
-  assume hyp: "(\<And>n\<Colon>inat. (\<And>m\<Colon>inat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
-  show "P n" by (blast intro: inat_less_induct hyp)
-qed
-
-subsection {* Complete Lattice *}
-
-instantiation inat :: complete_lattice
-begin
-
-definition inf_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
-  "inf_inat \<equiv> min"
-
-definition sup_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
-  "sup_inat \<equiv> max"
-
-definition Inf_inat :: "inat set \<Rightarrow> inat" where
-  "Inf_inat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
-
-definition Sup_inat :: "inat set \<Rightarrow> inat" where
-  "Sup_inat A \<equiv> if A = {} then 0
-    else if finite A then Max A
-                     else \<infinity>"
-instance proof
-  fix x :: "inat" and A :: "inat set"
-  { assume "x \<in> A" then show "Inf A \<le> x"
-      unfolding Inf_inat_def by (auto intro: Least_le) }
-  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
-      unfolding Inf_inat_def
-      by (cases "A = {}") (auto intro: LeastI2_ex) }
-  { assume "x \<in> A" then show "x \<le> Sup A"
-      unfolding Sup_inat_def by (cases "finite A") auto }
-  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
-      unfolding Sup_inat_def using finite_Fin_bounded by auto }
-qed (simp_all add: inf_inat_def sup_inat_def)
-end
-
-
-subsection {* Traditional theorem names *}
-
-lemmas inat_defs = zero_inat_def one_inat_def number_of_inat_def iSuc_def
-  plus_inat_def less_eq_inat_def less_inat_def
-
-end
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -8,87 +8,87 @@
 header {* Limits on the Extended real number line *}
 
 theory Extended_Real_Limits
-  imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Reals"
+  imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Real"
 begin
 
-lemma continuous_on_extreal[intro, simp]: "continuous_on A extreal"
-  unfolding continuous_on_topological open_extreal_def by auto
+lemma continuous_on_ereal[intro, simp]: "continuous_on A ereal"
+  unfolding continuous_on_topological open_ereal_def by auto
 
-lemma continuous_at_extreal[intro, simp]: "continuous (at x) extreal"
+lemma continuous_at_ereal[intro, simp]: "continuous (at x) ereal"
   using continuous_on_eq_continuous_at[of UNIV] by auto
 
-lemma continuous_within_extreal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) extreal"
+lemma continuous_within_ereal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) ereal"
   using continuous_on_eq_continuous_within[of A] by auto
 
-lemma extreal_open_uminus:
-  fixes S :: "extreal set"
+lemma ereal_open_uminus:
+  fixes S :: "ereal set"
   assumes "open S"
   shows "open (uminus ` S)"
-  unfolding open_extreal_def
+  unfolding open_ereal_def
 proof (intro conjI impI)
-  obtain x y where S: "open (extreal -` S)"
-    "\<infinity> \<in> S \<Longrightarrow> {extreal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< extreal y} \<subseteq> S"
-    using `open S` unfolding open_extreal_def by auto
-  have "extreal -` uminus ` S = uminus ` (extreal -` S)"
+  obtain x y where S: "open (ereal -` S)"
+    "\<infinity> \<in> S \<Longrightarrow> {ereal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< ereal y} \<subseteq> S"
+    using `open S` unfolding open_ereal_def by auto
+  have "ereal -` uminus ` S = uminus ` (ereal -` S)"
   proof safe
-    fix x y assume "extreal x = - y" "y \<in> S"
-    then show "x \<in> uminus ` extreal -` S" by (cases y) auto
+    fix x y assume "ereal x = - y" "y \<in> S"
+    then show "x \<in> uminus ` ereal -` S" by (cases y) auto
   next
-    fix x assume "extreal x \<in> S"
-    then show "- x \<in> extreal -` uminus ` S"
-      by (auto intro: image_eqI[of _ _ "extreal x"])
+    fix x assume "ereal x \<in> S"
+    then show "- x \<in> ereal -` uminus ` S"
+      by (auto intro: image_eqI[of _ _ "ereal x"])
   qed
-  then show "open (extreal -` uminus ` S)"
+  then show "open (ereal -` uminus ` S)"
     using S by (auto intro: open_negations)
   { assume "\<infinity> \<in> uminus ` S"
-    then have "-\<infinity> \<in> S" by (metis image_iff extreal_uminus_uminus)
-    then have "uminus ` {..<extreal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
-    then show "\<exists>x. {extreal x<..} \<subseteq> uminus ` S" using extreal_uminus_lessThan by auto }
+    then have "-\<infinity> \<in> S" by (metis image_iff ereal_uminus_uminus)
+    then have "uminus ` {..<ereal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
+    then show "\<exists>x. {ereal x<..} \<subseteq> uminus ` S" using ereal_uminus_lessThan by auto }
   { assume "-\<infinity> \<in> uminus ` S"
-    then have "\<infinity> : S" by (metis image_iff extreal_uminus_uminus)
-    then have "uminus ` {extreal x<..} <= uminus ` S" using S by (intro image_mono) auto
-    then show "\<exists>y. {..<extreal y} <= uminus ` S" using extreal_uminus_greaterThan by auto }
+    then have "\<infinity> : S" by (metis image_iff ereal_uminus_uminus)
+    then have "uminus ` {ereal x<..} <= uminus ` S" using S by (intro image_mono) auto
+    then show "\<exists>y. {..<ereal y} <= uminus ` S" using ereal_uminus_greaterThan by auto }
 qed
 
-lemma extreal_uminus_complement:
-  fixes S :: "extreal set"
+lemma ereal_uminus_complement:
+  fixes S :: "ereal set"
   shows "uminus ` (- S) = - uminus ` S"
   by (auto intro!: bij_image_Compl_eq surjI[of _ uminus] simp: bij_betw_def)
 
-lemma extreal_closed_uminus:
-  fixes S :: "extreal set"
+lemma ereal_closed_uminus:
+  fixes S :: "ereal set"
   assumes "closed S"
   shows "closed (uminus ` S)"
 using assms unfolding closed_def
-using extreal_open_uminus[of "- S"] extreal_uminus_complement by auto
+using ereal_open_uminus[of "- S"] ereal_uminus_complement by auto
 
-lemma not_open_extreal_singleton:
-  "\<not> (open {a :: extreal})"
+lemma not_open_ereal_singleton:
+  "\<not> (open {a :: ereal})"
 proof(rule ccontr)
   assume "\<not> \<not> open {a}" hence a: "open {a}" by auto
   show False
   proof (cases a)
     case MInf
-    then obtain y where "{..<extreal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
-    hence "extreal(y - 1):{a}" apply (subst subsetD[of "{..<extreal y}"]) by auto
+    then obtain y where "{..<ereal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
+    hence "ereal(y - 1):{a}" apply (subst subsetD[of "{..<ereal y}"]) by auto
     then show False using `a=(-\<infinity>)` by auto
   next
     case PInf
-    then obtain y where "{extreal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
-    hence "extreal(y+1):{a}" apply (subst subsetD[of "{extreal y<..}"]) by auto
+    then obtain y where "{ereal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
+    hence "ereal(y+1):{a}" apply (subst subsetD[of "{ereal y<..}"]) by auto
     then show False using `a=\<infinity>` by auto
   next
     case (real r) then have fin: "\<bar>a\<bar> \<noteq> \<infinity>" by simp
-    from extreal_open_cont_interval[OF a singletonI this] guess e . note e = this
+    from ereal_open_cont_interval[OF a singletonI this] guess e . note e = this
     then obtain b where b_def: "a<b & b<a+e"
-      using fin extreal_between extreal_dense[of a "a+e"] by auto
-    then have "b: {a-e <..< a+e}" using fin extreal_between[of a e] e by auto
+      using fin ereal_between ereal_dense[of a "a+e"] by auto
+    then have "b: {a-e <..< a+e}" using fin ereal_between[of a e] e by auto
     then show False using b_def e by auto
   qed
 qed
 
-lemma extreal_closed_contains_Inf:
-  fixes S :: "extreal set"
+lemma ereal_closed_contains_Inf:
+  fixes S :: "ereal set"
   assumes "closed S" "S ~= {}"
   shows "Inf S : S"
 proof(rule ccontr)
@@ -96,8 +96,8 @@
   show False
   proof (cases "Inf S")
     case MInf hence "(-\<infinity>) : - S" using a by auto
-    then obtain y where "{..<extreal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
-    hence "extreal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff
+    then obtain y where "{..<ereal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
+    hence "ereal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff
       complete_lattice_class.Inf_greatest double_complement set_rev_mp)
     then show False using MInf by auto
   next
@@ -105,9 +105,9 @@
     then show False by (metis `Inf S ~: S` insert_code mem_def PInf)
   next
     case (real r) then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp
-    from extreal_open_cont_interval[OF a this] guess e . note e = this
+    from ereal_open_cont_interval[OF a this] guess e . note e = this
     { fix x assume "x:S" hence "x>=Inf S" by (rule complete_lattice_class.Inf_lower)
-      hence *: "x>Inf S-e" using e by (metis fin extreal_between(1) order_less_le_trans)
+      hence *: "x>Inf S-e" using e by (metis fin ereal_between(1) order_less_le_trans)
       { assume "x<Inf S+e" hence "x:{Inf S-e <..< Inf S+e}" using * by auto
         hence False using e `x:S` by auto
       } hence "x>=Inf S+e" by (metis linorder_le_less_linear)
@@ -116,115 +116,117 @@
   qed
 qed
 
-lemma extreal_closed_contains_Sup:
-  fixes S :: "extreal set"
+lemma ereal_closed_contains_Sup:
+  fixes S :: "ereal set"
   assumes "closed S" "S ~= {}"
   shows "Sup S : S"
 proof-
-  have "closed (uminus ` S)" by (metis assms(1) extreal_closed_uminus)
-  hence "Inf (uminus ` S) : uminus ` S" using assms extreal_closed_contains_Inf[of "uminus ` S"] by auto
-  hence "- Sup S : uminus ` S" using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image)
-  thus ?thesis by (metis imageI extreal_uminus_uminus extreal_minus_minus_image)
+  have "closed (uminus ` S)" by (metis assms(1) ereal_closed_uminus)
+  hence "Inf (uminus ` S) : uminus ` S" using assms ereal_closed_contains_Inf[of "uminus ` S"] by auto
+  hence "- Sup S : uminus ` S" using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image)
+  thus ?thesis by (metis imageI ereal_uminus_uminus ereal_minus_minus_image)
 qed
 
-lemma extreal_open_closed_aux:
-  fixes S :: "extreal set"
+lemma ereal_open_closed_aux:
+  fixes S :: "ereal set"
   assumes "open S" "closed S"
   assumes S: "(-\<infinity>) ~: S"
   shows "S = {}"
 proof(rule ccontr)
   assume "S ~= {}"
-  hence *: "(Inf S):S" by (metis assms(2) extreal_closed_contains_Inf)
+  hence *: "(Inf S):S" by (metis assms(2) ereal_closed_contains_Inf)
   { assume "Inf S=(-\<infinity>)" hence False using * assms(3) by auto }
   moreover
   { assume "Inf S=\<infinity>" hence "S={\<infinity>}" by (metis Inf_eq_PInfty `S ~= {}`)
-    hence False by (metis assms(1) not_open_extreal_singleton) }
+    hence False by (metis assms(1) not_open_ereal_singleton) }
   moreover
   { assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"
-    from extreal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
+    from ereal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
     then obtain b where b_def: "Inf S-e<b & b<Inf S"
-      using fin extreal_between[of "Inf S" e] extreal_dense[of "Inf S-e"] by auto
-    hence "b: {Inf S-e <..< Inf S+e}" using e fin extreal_between[of "Inf S" e] by auto
+      using fin ereal_between[of "Inf S" e] ereal_dense[of "Inf S-e"] by auto
+    hence "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e] by auto
     hence "b:S" using e by auto
     hence False using b_def by (metis complete_lattice_class.Inf_lower leD)
   } ultimately show False by auto
 qed
 
-lemma extreal_open_closed:
-  fixes S :: "extreal set"
+lemma ereal_open_closed:
+  fixes S :: "ereal set"
   shows "(open S & closed S) <-> (S = {} | S = UNIV)"
 proof-
 { assume lhs: "open S & closed S"
-  { assume "(-\<infinity>) ~: S" hence "S={}" using lhs extreal_open_closed_aux by auto }
+  { assume "(-\<infinity>) ~: S" hence "S={}" using lhs ereal_open_closed_aux by auto }
   moreover
-  { assume "(-\<infinity>) : S" hence "(- S)={}" using lhs extreal_open_closed_aux[of "-S"] by auto }
+  { assume "(-\<infinity>) : S" hence "(- S)={}" using lhs ereal_open_closed_aux[of "-S"] by auto }
   ultimately have "S = {} | S = UNIV" by auto
 } thus ?thesis by auto
 qed
 
-lemma extreal_open_affinity_pos:
+lemma ereal_open_affinity_pos:
+  fixes S :: "ereal set"
   assumes "open S" and m: "m \<noteq> \<infinity>" "0 < m" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
   shows "open ((\<lambda>x. m * x + t) ` S)"
 proof -
-  obtain r where r[simp]: "m = extreal r" using m by (cases m) auto
-  obtain p where p[simp]: "t = extreal p" using t by auto
+  obtain r where r[simp]: "m = ereal r" using m by (cases m) auto
+  obtain p where p[simp]: "t = ereal p" using t by auto
   have "r \<noteq> 0" "0 < r" and m': "m \<noteq> \<infinity>" "m \<noteq> -\<infinity>" "m \<noteq> 0" using m by auto
-  from `open S`[THEN extreal_openE] guess l u . note T = this
+  from `open S`[THEN ereal_openE] guess l u . note T = this
   let ?f = "(\<lambda>x. m * x + t)"
-  show ?thesis unfolding open_extreal_def
+  show ?thesis unfolding open_ereal_def
   proof (intro conjI impI exI subsetI)
-    have "extreal -` ?f ` S = (\<lambda>x. r * x + p) ` (extreal -` S)"
+    have "ereal -` ?f ` S = (\<lambda>x. r * x + p) ` (ereal -` S)"
     proof safe
-      fix x y assume "extreal y = m * x + t" "x \<in> S"
-      then show "y \<in> (\<lambda>x. r * x + p) ` extreal -` S"
+      fix x y assume "ereal y = m * x + t" "x \<in> S"
+      then show "y \<in> (\<lambda>x. r * x + p) ` ereal -` S"
         using `r \<noteq> 0` by (cases x) (auto intro!: image_eqI[of _ _ "real x"] split: split_if_asm)
     qed force
-    then show "open (extreal -` ?f ` S)"
+    then show "open (ereal -` ?f ` S)"
       using open_affinity[OF T(1) `r \<noteq> 0`] by (auto simp: ac_simps)
   next
     assume "\<infinity> \<in> ?f`S" with `0 < r` have "\<infinity> \<in> S" by auto
-    fix x assume "x \<in> {extreal (r * l + p)<..}"
-    then have [simp]: "extreal (r * l + p) < x" by auto
+    fix x assume "x \<in> {ereal (r * l + p)<..}"
+    then have [simp]: "ereal (r * l + p) < x" by auto
     show "x \<in> ?f`S"
     proof (rule image_eqI)
       show "x = m * ((x - t) / m) + t"
-        using m t by (cases rule: extreal3_cases[of m x t]) auto
-      have "extreal l < (x - t)/m"
-        using m t by (simp add: extreal_less_divide_pos extreal_less_minus)
+        using m t by (cases rule: ereal3_cases[of m x t]) auto
+      have "ereal l < (x - t)/m"
+        using m t by (simp add: ereal_less_divide_pos ereal_less_minus)
       then show "(x - t)/m \<in> S" using T(2)[OF `\<infinity> \<in> S`] by auto
     qed
   next
     assume "-\<infinity> \<in> ?f`S" with `0 < r` have "-\<infinity> \<in> S" by auto
-    fix x assume "x \<in> {..<extreal (r * u + p)}"
-    then have [simp]: "x < extreal (r * u + p)" by auto
+    fix x assume "x \<in> {..<ereal (r * u + p)}"
+    then have [simp]: "x < ereal (r * u + p)" by auto
     show "x \<in> ?f`S"
     proof (rule image_eqI)
       show "x = m * ((x - t) / m) + t"
-        using m t by (cases rule: extreal3_cases[of m x t]) auto
-      have "(x - t)/m < extreal u"
-        using m t by (simp add: extreal_divide_less_pos extreal_minus_less)
+        using m t by (cases rule: ereal3_cases[of m x t]) auto
+      have "(x - t)/m < ereal u"
+        using m t by (simp add: ereal_divide_less_pos ereal_minus_less)
       then show "(x - t)/m \<in> S" using T(3)[OF `-\<infinity> \<in> S`] by auto
     qed
   qed
 qed
 
-lemma extreal_open_affinity:
+lemma ereal_open_affinity:
+  fixes S :: "ereal set"
   assumes "open S" and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
   shows "open ((\<lambda>x. m * x + t) ` S)"
 proof cases
   assume "0 < m" then show ?thesis
-    using extreal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto
+    using ereal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto
 next
   assume "\<not> 0 < m" then
   have "0 < -m" using `m \<noteq> 0` by (cases m) auto
   then have m: "-m \<noteq> \<infinity>" "0 < -m" using `\<bar>m\<bar> \<noteq> \<infinity>`
-    by (auto simp: extreal_uminus_eq_reorder)
-  from extreal_open_affinity_pos[OF extreal_open_uminus[OF `open S`] m t]
+    by (auto simp: ereal_uminus_eq_reorder)
+  from ereal_open_affinity_pos[OF ereal_open_uminus[OF `open S`] m t]
   show ?thesis unfolding image_image by simp
 qed
 
-lemma extreal_lim_mult:
-  fixes X :: "'a \<Rightarrow> extreal"
+lemma ereal_lim_mult:
+  fixes X :: "'a \<Rightarrow> ereal"
   assumes lim: "(X ---> L) net" and a: "\<bar>a\<bar> \<noteq> \<infinity>"
   shows "((\<lambda>i. a * X i) ---> a * L) net"
 proof cases
@@ -233,73 +235,73 @@
   proof (rule topological_tendstoI)
     fix S assume "open S" "a * L \<in> S"
     have "a * L / a = L"
-      using `a \<noteq> 0` a by (cases rule: extreal2_cases[of a L]) auto
+      using `a \<noteq> 0` a by (cases rule: ereal2_cases[of a L]) auto
     then have L: "L \<in> ((\<lambda>x. x / a) ` S)"
       using `a * L \<in> S` by (force simp: image_iff)
     moreover have "open ((\<lambda>x. x / a) ` S)"
-      using extreal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a
-      by (auto simp: extreal_divide_eq extreal_inverse_eq_0 divide_extreal_def ac_simps)
+      using ereal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a
+      by (auto simp: ereal_divide_eq ereal_inverse_eq_0 divide_ereal_def ac_simps)
     note * = lim[THEN topological_tendstoD, OF this L]
     { fix x from a `a \<noteq> 0` have "a * (x / a) = x"
-        by (cases rule: extreal2_cases[of a x]) auto }
+        by (cases rule: ereal2_cases[of a x]) auto }
     note this[simp]
     show "eventually (\<lambda>x. a * X x \<in> S) net"
       by (rule eventually_mono[OF _ *]) auto
   qed
 qed auto
 
-lemma extreal_lim_uminus:
-  fixes X :: "'a \<Rightarrow> extreal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
-  using extreal_lim_mult[of X L net "extreal (-1)"]
-        extreal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "extreal (-1)"]
+lemma ereal_lim_uminus:
+  fixes X :: "'a \<Rightarrow> ereal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
+  using ereal_lim_mult[of X L net "ereal (-1)"]
+        ereal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "ereal (-1)"]
   by (auto simp add: algebra_simps)
 
-lemma Lim_bounded2_extreal:
-  assumes lim:"f ----> (l :: extreal)"
+lemma Lim_bounded2_ereal:
+  assumes lim:"f ----> (l :: ereal)"
   and ge: "ALL n>=N. f n >= C"
   shows "l>=C"
 proof-
 def g == "(%i. -(f i))"
-{ fix n assume "n>=N" hence "g n <= -C" using assms extreal_minus_le_minus g_def by auto }
+{ fix n assume "n>=N" hence "g n <= -C" using assms ereal_minus_le_minus g_def by auto }
 hence "ALL n>=N. g n <= -C" by auto
-moreover have limg: "g ----> (-l)" using g_def extreal_lim_uminus lim by auto
-ultimately have "-l <= -C" using Lim_bounded_extreal[of g "-l" _ "-C"] by auto
-from this show ?thesis using extreal_minus_le_minus by auto
+moreover have limg: "g ----> (-l)" using g_def ereal_lim_uminus lim by auto
+ultimately have "-l <= -C" using Lim_bounded_ereal[of g "-l" _ "-C"] by auto
+from this show ?thesis using ereal_minus_le_minus by auto
 qed
 
 
-lemma extreal_open_atLeast: "open {x..} \<longleftrightarrow> x = -\<infinity>"
+lemma ereal_open_atLeast: fixes x :: ereal shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
 proof
   assume "x = -\<infinity>" then have "{x..} = UNIV" by auto
   then show "open {x..}" by auto
 next
   assume "open {x..}"
   then have "open {x..} \<and> closed {x..}" by auto
-  then have "{x..} = UNIV" unfolding extreal_open_closed by auto
-  then show "x = -\<infinity>" by (simp add: bot_extreal_def atLeast_eq_UNIV_iff)
+  then have "{x..} = UNIV" unfolding ereal_open_closed by auto
+  then show "x = -\<infinity>" by (simp add: bot_ereal_def atLeast_eq_UNIV_iff)
 qed
 
-lemma extreal_open_mono_set:
-  fixes S :: "extreal set"
+lemma ereal_open_mono_set:
+  fixes S :: "ereal set"
   defines "a \<equiv> Inf S"
   shows "(open S \<and> mono S) \<longleftrightarrow> (S = UNIV \<or> S = {a <..})"
-  by (metis Inf_UNIV a_def atLeast_eq_UNIV_iff extreal_open_atLeast
-            extreal_open_closed mono_set_iff open_extreal_greaterThan)
+  by (metis Inf_UNIV a_def atLeast_eq_UNIV_iff ereal_open_atLeast
+            ereal_open_closed mono_set_iff open_ereal_greaterThan)
 
-lemma extreal_closed_mono_set:
-  fixes S :: "extreal set"
+lemma ereal_closed_mono_set:
+  fixes S :: "ereal set"
   shows "(closed S \<and> mono S) \<longleftrightarrow> (S = {} \<or> S = {Inf S ..})"
-  by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_extreal_atLeast
-            extreal_open_closed mono_empty mono_set_iff open_extreal_greaterThan)
+  by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast
+            ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan)
 
-lemma extreal_Liminf_Sup_monoset:
-  fixes f :: "'a => extreal"
+lemma ereal_Liminf_Sup_monoset:
+  fixes f :: "'a => ereal"
   shows "Liminf net f = Sup {l. \<forall>S. open S \<longrightarrow> mono S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
   unfolding Liminf_Sup
 proof (intro arg_cong[where f="\<lambda>P. Sup (Collect P)"] ext iffI allI impI)
   fix l S assume ev: "\<forall>y<l. eventually (\<lambda>x. y < f x) net" and "open S" "mono S" "l \<in> S"
   then have "S = UNIV \<or> S = {Inf S <..}"
-    using extreal_open_mono_set[of S] by auto
+    using ereal_open_mono_set[of S] by auto
   then show "eventually (\<lambda>x. f x \<in> S) net"
   proof
     assume S: "S = {Inf S<..}"
@@ -314,15 +316,15 @@
   then show "eventually (\<lambda>x. y < f x) net" by auto
 qed
 
-lemma extreal_Limsup_Inf_monoset:
-  fixes f :: "'a => extreal"
+lemma ereal_Limsup_Inf_monoset:
+  fixes f :: "'a => ereal"
   shows "Limsup net f = Inf {l. \<forall>S. open S \<longrightarrow> mono (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
   unfolding Limsup_Inf
 proof (intro arg_cong[where f="\<lambda>P. Inf (Collect P)"] ext iffI allI impI)
   fix l S assume ev: "\<forall>y>l. eventually (\<lambda>x. f x < y) net" and "open S" "mono (uminus`S)" "l \<in> S"
-  then have "open (uminus`S) \<and> mono (uminus`S)" by (simp add: extreal_open_uminus)
+  then have "open (uminus`S) \<and> mono (uminus`S)" by (simp add: ereal_open_uminus)
   then have "S = UNIV \<or> S = {..< Sup S}"
-    unfolding extreal_open_mono_set extreal_Inf_uminus_image_eq extreal_image_uminus_shift by simp
+    unfolding ereal_open_mono_set ereal_Inf_uminus_image_eq ereal_image_uminus_shift by simp
   then show "eventually (\<lambda>x. f x \<in> S) net"
   proof
     assume S: "S = {..< Sup S}"
@@ -338,70 +340,70 @@
 qed
 
 
-lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::extreal set)"
-  using extreal_open_uminus[of S] extreal_open_uminus[of "uminus`S"] by auto
+lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::ereal set)"
+  using ereal_open_uminus[of S] ereal_open_uminus[of "uminus`S"] by auto
 
-lemma extreal_Limsup_uminus:
-  fixes f :: "'a => extreal"
+lemma ereal_Limsup_uminus:
+  fixes f :: "'a => ereal"
   shows "Limsup net (\<lambda>x. - (f x)) = -(Liminf net f)"
 proof -
-  { fix P l have "(\<exists>x. (l::extreal) = -x \<and> P x) \<longleftrightarrow> P (-l)" by (auto intro!: exI[of _ "-l"]) }
+  { fix P l have "(\<exists>x. (l::ereal) = -x \<and> P x) \<longleftrightarrow> P (-l)" by (auto intro!: exI[of _ "-l"]) }
   note Ex_cancel = this
-  { fix P :: "extreal set \<Rightarrow> bool" have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))"
+  { fix P :: "ereal set \<Rightarrow> bool" have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))"
       apply auto by (erule_tac x="uminus`S" in allE) (auto simp: image_image) }
   note add_uminus_image = this
-  { fix x S have "(x::extreal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" by (auto intro!: image_eqI[of _ _ "-x"]) }
+  { fix x S have "(x::ereal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" by (auto intro!: image_eqI[of _ _ "-x"]) }
   note remove_uminus_image = this
   show ?thesis
-    unfolding extreal_Limsup_Inf_monoset extreal_Liminf_Sup_monoset
-    unfolding extreal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel
+    unfolding ereal_Limsup_Inf_monoset ereal_Liminf_Sup_monoset
+    unfolding ereal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel
     by (subst add_uminus_image) (simp add: open_uminus_iff remove_uminus_image)
 qed
 
-lemma extreal_Liminf_uminus:
-  fixes f :: "'a => extreal"
+lemma ereal_Liminf_uminus:
+  fixes f :: "'a => ereal"
   shows "Liminf net (\<lambda>x. - (f x)) = -(Limsup net f)"
-  using extreal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto
+  using ereal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto
 
-lemma extreal_Lim_uminus:
-  fixes f :: "'a \<Rightarrow> extreal" shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net"
+lemma ereal_Lim_uminus:
+  fixes f :: "'a \<Rightarrow> ereal" shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net"
   using
-    extreal_lim_mult[of f f0 net "- 1"]
-    extreal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"]
-  by (auto simp: extreal_uminus_reorder)
+    ereal_lim_mult[of f f0 net "- 1"]
+    ereal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"]
+  by (auto simp: ereal_uminus_reorder)
 
 lemma lim_imp_Limsup:
-  fixes f :: "'a => extreal"
+  fixes f :: "'a => ereal"
   assumes "\<not> trivial_limit net"
   assumes lim: "(f ---> f0) net"
   shows "Limsup net f = f0"
-  using extreal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"]
-     extreal_Liminf_uminus[of net f] assms by simp
+  using ereal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"]
+     ereal_Liminf_uminus[of net f] assms by simp
 
 lemma Liminf_PInfty:
-  fixes f :: "'a \<Rightarrow> extreal"
+  fixes f :: "'a \<Rightarrow> ereal"
   assumes "\<not> trivial_limit net"
   shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>"
 proof (intro lim_imp_Liminf iffI assms)
   assume rhs: "Liminf net f = \<infinity>"
-  { fix S assume "open S & \<infinity> : S"
-    then obtain m where "{extreal m<..} <= S" using open_PInfty2 by auto
-    moreover have "eventually (\<lambda>x. f x \<in> {extreal m<..}) net"
-      using rhs unfolding Liminf_Sup top_extreal_def[symmetric] Sup_eq_top_iff
-      by (auto elim!: allE[where x="extreal m"] simp: top_extreal_def)
+  { fix S :: "ereal set" assume "open S & \<infinity> : S"
+    then obtain m where "{ereal m<..} <= S" using open_PInfty2 by auto
+    moreover have "eventually (\<lambda>x. f x \<in> {ereal m<..}) net"
+      using rhs unfolding Liminf_Sup top_ereal_def[symmetric] Sup_eq_top_iff
+      by (auto elim!: allE[where x="ereal m"] simp: top_ereal_def)
     ultimately have "eventually (%x. f x : S) net" apply (subst eventually_mono) by auto
   } then show "(f ---> \<infinity>) net" unfolding tendsto_def by auto
 qed
 
 lemma Limsup_MInfty:
-  fixes f :: "'a \<Rightarrow> extreal"
+  fixes f :: "'a \<Rightarrow> ereal"
   assumes "\<not> trivial_limit net"
   shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>"
-  using assms extreal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"]
-        extreal_Liminf_uminus[of _ f] by (auto simp: extreal_uminus_eq_reorder)
+  using assms ereal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"]
+        ereal_Liminf_uminus[of _ f] by (auto simp: ereal_uminus_eq_reorder)
 
-lemma extreal_Liminf_eq_Limsup:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma ereal_Liminf_eq_Limsup:
+  fixes f :: "'a \<Rightarrow> ereal"
   assumes ntriv: "\<not> trivial_limit net"
   assumes lim: "Liminf net f = f0" "Limsup net f = f0"
   shows "(f ---> f0) net"
@@ -415,7 +417,7 @@
   proof (rule topological_tendstoI)
     fix S assume "open S""f0 \<in> S"
     then obtain a b where "a < Liminf net f" "Limsup net f < b" "{a<..<b} \<subseteq> S"
-      using extreal_open_cont_interval2[of S f0] real lim by auto
+      using ereal_open_cont_interval2[of S f0] real lim by auto
     then have "eventually (\<lambda>x. f x \<in> {a<..<b}) net"
       unfolding Liminf_Sup Limsup_Inf less_Sup_iff Inf_less_iff
       by (auto intro!: eventually_conj simp add: greaterThanLessThan_iff)
@@ -424,62 +426,62 @@
   qed
 qed
 
-lemma extreal_Liminf_eq_Limsup_iff:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma ereal_Liminf_eq_Limsup_iff:
+  fixes f :: "'a \<Rightarrow> ereal"
   assumes "\<not> trivial_limit net"
   shows "(f ---> f0) net \<longleftrightarrow> Liminf net f = f0 \<and> Limsup net f = f0"
-  by (metis assms extreal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup)
+  by (metis assms ereal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup)
 
 lemma limsup_INFI_SUPR:
-  fixes f :: "nat \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> ereal"
   shows "limsup f = (INF n. SUP m:{n..}. f m)"
-  using extreal_Limsup_uminus[of sequentially "\<lambda>x. - f x"]
-  by (simp add: liminf_SUPR_INFI extreal_INFI_uminus extreal_SUPR_uminus)
+  using ereal_Limsup_uminus[of sequentially "\<lambda>x. - f x"]
+  by (simp add: liminf_SUPR_INFI ereal_INFI_uminus ereal_SUPR_uminus)
 
 lemma liminf_PInfty:
-  fixes X :: "nat => extreal"
+  fixes X :: "nat => ereal"
   shows "X ----> \<infinity> <-> liminf X = \<infinity>"
 by (metis Liminf_PInfty trivial_limit_sequentially)
 
 lemma limsup_MInfty:
-  fixes X :: "nat => extreal"
+  fixes X :: "nat => ereal"
   shows "X ----> (-\<infinity>) <-> limsup X = (-\<infinity>)"
 by (metis Limsup_MInfty trivial_limit_sequentially)
 
-lemma extreal_lim_mono:
-  fixes X Y :: "nat => extreal"
+lemma ereal_lim_mono:
+  fixes X Y :: "nat => ereal"
   assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
   assumes "X ----> x" "Y ----> y"
   shows "x <= y"
-  by (metis extreal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
+  by (metis ereal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
 
-lemma incseq_le_extreal:
-  fixes X :: "nat \<Rightarrow> extreal"
+lemma incseq_le_ereal:
+  fixes X :: "nat \<Rightarrow> ereal"
   assumes inc: "incseq X" and lim: "X ----> L"
   shows "X N \<le> L"
   using inc
-  by (intro extreal_lim_mono[of N, OF _ Lim_const lim]) (simp add: incseq_def)
+  by (intro ereal_lim_mono[of N, OF _ Lim_const lim]) (simp add: incseq_def)
 
-lemma decseq_ge_extreal: assumes dec: "decseq X"
-  and lim: "X ----> (L::extreal)" shows "X N >= L"
+lemma decseq_ge_ereal: assumes dec: "decseq X"
+  and lim: "X ----> (L::ereal)" shows "X N >= L"
   using dec
-  by (intro extreal_lim_mono[of N, OF _ lim Lim_const]) (simp add: decseq_def)
+  by (intro ereal_lim_mono[of N, OF _ lim Lim_const]) (simp add: decseq_def)
 
 lemma liminf_bounded_open:
-  fixes x :: "nat \<Rightarrow> extreal"
+  fixes x :: "nat \<Rightarrow> ereal"
   shows "x0 \<le> liminf x \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> mono S \<longrightarrow> x0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. x n \<in> S))" 
   (is "_ \<longleftrightarrow> ?P x0")
 proof
   assume "?P x0" then show "x0 \<le> liminf x"
-    unfolding extreal_Liminf_Sup_monoset eventually_sequentially
+    unfolding ereal_Liminf_Sup_monoset eventually_sequentially
     by (intro complete_lattice_class.Sup_upper) auto
 next
   assume "x0 \<le> liminf x"
-  { fix S :: "extreal set" assume om: "open S & mono S & x0:S"
+  { fix S :: "ereal set" assume om: "open S & mono S & x0:S"
     { assume "S = UNIV" hence "EX N. (ALL n>=N. x n : S)" by auto }
     moreover
     { assume "~(S=UNIV)"
-      then obtain B where B_def: "S = {B<..}" using om extreal_open_mono_set by auto
+      then obtain B where B_def: "S = {B<..}" using om ereal_open_mono_set by auto
       hence "B<x0" using om by auto
       hence "EX N. ALL n>=N. x n : S" unfolding B_def using `x0 \<le> liminf x` liminf_bounded_iff by auto
     } ultimately have "EX N. (ALL n>=N. x n : S)" by auto
@@ -487,15 +489,15 @@
 qed
 
 lemma limsup_subseq_mono:
-  fixes X :: "nat \<Rightarrow> extreal"
+  fixes X :: "nat \<Rightarrow> ereal"
   assumes "subseq r"
   shows "limsup (X \<circ> r) \<le> limsup X"
 proof-
   have "(\<lambda>n. - X n) \<circ> r = (\<lambda>n. - (X \<circ> r) n)" by (simp add: fun_eq_iff)
   then have "- limsup X \<le> - limsup (X \<circ> r)"
      using liminf_subseq_mono[of r "(%n. - X n)"]
-       extreal_Liminf_uminus[of sequentially X]
-       extreal_Liminf_uminus[of sequentially "X o r"] assms by auto
+       ereal_Liminf_uminus[of sequentially X]
+       ereal_Liminf_uminus[of sequentially "X o r"] assms by auto
   then show ?thesis by auto
 qed
 
@@ -514,8 +516,8 @@
 from this show ?thesis apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
    using assms by auto
 qed
-lemma lim_extreal_increasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
-  obtains l where "f ----> (l::extreal)"
+lemma lim_ereal_increasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
+  obtains l where "f ----> (l::ereal)"
 proof(cases "f = (\<lambda>x. - \<infinity>)")
   case True then show thesis using Lim_const[of "- \<infinity>" sequentially] by (intro that[of "-\<infinity>"]) auto
 next
@@ -527,18 +529,18 @@
   hence incy: "!!n m. n>=m ==> Y n >= Y m" using assms by auto
   from minf have minfy: "ALL n. Y n ~= (-\<infinity>)" using Y_def by auto
   show thesis
-  proof(cases "EX B. ALL n. f n < extreal B")
+  proof(cases "EX B. ALL n. f n < ereal B")
     case False thus thesis apply- apply(rule that[of \<infinity>]) unfolding Lim_PInfty not_ex not_all
     apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
     apply(rule order_trans[OF _ assms[rule_format]]) by auto
   next case True then guess B ..
-    hence "ALL n. Y n < extreal B" using Y_def by auto note B = this[rule_format]
+    hence "ALL n. Y n < ereal B" using Y_def by auto note B = this[rule_format]
     { fix n have "Y n < \<infinity>" using B[of n] apply (subst less_le_trans) by auto
       hence "Y n ~= \<infinity> & Y n ~= (-\<infinity>)" using minfy by auto
     } hence *: "ALL n. \<bar>Y n\<bar> \<noteq> \<infinity>" by auto
     { fix n have "real (Y n) < B" proof- case goal1 thus ?case
-        using B[of n] apply-apply(subst(asm) extreal_real'[THEN sym]) defer defer
-        unfolding extreal_less using * by auto
+        using B[of n] apply-apply(subst(asm) ereal_real'[THEN sym]) defer defer
+        unfolding ereal_less using * by auto
       qed
     }
     hence B': "ALL n. (real (Y n) <= B)" using less_imp_le by auto
@@ -546,29 +548,29 @@
       apply(rule bounded_increasing_convergent2)
     proof safe show "!!n. real (Y n) <= B" using B' by auto
       fix n m::nat assume "n<=m"
-      hence "extreal (real (Y n)) <= extreal (real (Y m))"
-        using incy[rule_format,of n m] apply(subst extreal_real)+
+      hence "ereal (real (Y n)) <= ereal (real (Y m))"
+        using incy[rule_format,of n m] apply(subst ereal_real)+
         using *[rule_format, of n] *[rule_format, of m] by auto
       thus "real (Y n) <= real (Y m)" by auto
     qed then guess l .. note l=this
-    have "Y ----> extreal l" using l apply-apply(subst(asm) lim_extreal[THEN sym])
-    unfolding extreal_real using * by auto
-    thus thesis apply-apply(rule that[of "extreal l"])
+    have "Y ----> ereal l" using l apply-apply(subst(asm) lim_ereal[THEN sym])
+    unfolding ereal_real using * by auto
+    thus thesis apply-apply(rule that[of "ereal l"])
        apply (subst tail_same_limit[of Y _ N]) using Y_def by auto
   qed
 qed
 
-lemma lim_extreal_decreasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
-  obtains l where "f ----> (l::extreal)"
+lemma lim_ereal_decreasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
+  obtains l where "f ----> (l::ereal)"
 proof -
-  from lim_extreal_increasing[of "\<lambda>x. - f x"] assms
+  from lim_ereal_increasing[of "\<lambda>x. - f x"] assms
   obtain l where "(\<lambda>x. - f x) ----> l" by auto
-  from extreal_lim_mult[OF this, of "- 1"] show thesis
-    by (intro that[of "-l"]) (simp add: extreal_uminus_eq_reorder)
+  from ereal_lim_mult[OF this, of "- 1"] show thesis
+    by (intro that[of "-l"]) (simp add: ereal_uminus_eq_reorder)
 qed
 
-lemma compact_extreal:
-  fixes X :: "nat \<Rightarrow> extreal"
+lemma compact_ereal:
+  fixes X :: "nat \<Rightarrow> ereal"
   shows "\<exists>l r. subseq r \<and> (X \<circ> r) ----> l"
 proof -
   obtain r where "subseq r" and mono: "monoseq (X \<circ> r)"
@@ -576,66 +578,66 @@
   then have "(\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) m \<le> (X \<circ> r) n) \<or> (\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) n \<le> (X \<circ> r) m)"
     by (auto simp add: monoseq_def)
   then obtain l where "(X\<circ>r) ----> l"
-     using lim_extreal_increasing[of "X \<circ> r"] lim_extreal_decreasing[of "X \<circ> r"] by auto
+     using lim_ereal_increasing[of "X \<circ> r"] lim_ereal_decreasing[of "X \<circ> r"] by auto
   then show ?thesis using `subseq r` by auto
 qed
 
-lemma extreal_Sup_lim:
-  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
+lemma ereal_Sup_lim:
+  assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)"
   shows "a \<le> Sup s"
-by (metis Lim_bounded_extreal assms complete_lattice_class.Sup_upper)
+by (metis Lim_bounded_ereal assms complete_lattice_class.Sup_upper)
 
-lemma extreal_Inf_lim:
-  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
+lemma ereal_Inf_lim:
+  assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)"
   shows "Inf s \<le> a"
-by (metis Lim_bounded2_extreal assms complete_lattice_class.Inf_lower)
+by (metis Lim_bounded2_ereal assms complete_lattice_class.Inf_lower)
 
-lemma SUP_Lim_extreal:
-  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" "X ----> l" shows "(SUP n. X n) = l"
-proof (rule extreal_SUPI)
+lemma SUP_Lim_ereal:
+  fixes X :: "nat \<Rightarrow> ereal" assumes "incseq X" "X ----> l" shows "(SUP n. X n) = l"
+proof (rule ereal_SUPI)
   fix n from assms show "X n \<le> l"
-    by (intro incseq_le_extreal) (simp add: incseq_def)
+    by (intro incseq_le_ereal) (simp add: incseq_def)
 next
   fix y assume "\<And>n. n \<in> UNIV \<Longrightarrow> X n \<le> y"
-  with extreal_Sup_lim[OF _ `X ----> l`, of "{..y}"]
+  with ereal_Sup_lim[OF _ `X ----> l`, of "{..y}"]
   show "l \<le> y" by auto
 qed
 
-lemma LIMSEQ_extreal_SUPR:
-  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" shows "X ----> (SUP n. X n)"
-proof (rule lim_extreal_increasing)
+lemma LIMSEQ_ereal_SUPR:
+  fixes X :: "nat \<Rightarrow> ereal" assumes "incseq X" shows "X ----> (SUP n. X n)"
+proof (rule lim_ereal_increasing)
   fix n m :: nat assume "m \<le> n" then show "X m \<le> X n"
     using `incseq X` by (simp add: incseq_def)
 next
   fix l assume "X ----> l"
-  with SUP_Lim_extreal[of X, OF assms this] show ?thesis by simp
+  with SUP_Lim_ereal[of X, OF assms this] show ?thesis by simp
 qed
 
-lemma INF_Lim_extreal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::extreal)"
-  using SUP_Lim_extreal[of "\<lambda>i. - X i" "- l"]
-  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
+lemma INF_Lim_ereal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::ereal)"
+  using SUP_Lim_ereal[of "\<lambda>i. - X i" "- l"]
+  by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
 
-lemma LIMSEQ_extreal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: extreal)"
-  using LIMSEQ_extreal_SUPR[of "\<lambda>i. - X i"]
-  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
+lemma LIMSEQ_ereal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: ereal)"
+  using LIMSEQ_ereal_SUPR[of "\<lambda>i. - X i"]
+  by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
 
 lemma SUP_eq_LIMSEQ:
   assumes "mono f"
-  shows "(SUP n. extreal (f n)) = extreal x \<longleftrightarrow> f ----> x"
+  shows "(SUP n. ereal (f n)) = ereal x \<longleftrightarrow> f ----> x"
 proof
-  have inc: "incseq (\<lambda>i. extreal (f i))"
+  have inc: "incseq (\<lambda>i. ereal (f i))"
     using `mono f` unfolding mono_def incseq_def by auto
   { assume "f ----> x"
-   then have "(\<lambda>i. extreal (f i)) ----> extreal x" by auto
-   from SUP_Lim_extreal[OF inc this]
-   show "(SUP n. extreal (f n)) = extreal x" . }
-  { assume "(SUP n. extreal (f n)) = extreal x"
-    with LIMSEQ_extreal_SUPR[OF inc]
+   then have "(\<lambda>i. ereal (f i)) ----> ereal x" by auto
+   from SUP_Lim_ereal[OF inc this]
+   show "(SUP n. ereal (f n)) = ereal x" . }
+  { assume "(SUP n. ereal (f n)) = ereal x"
+    with LIMSEQ_ereal_SUPR[OF inc]
     show "f ----> x" by auto }
 qed
 
 lemma Liminf_within:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
 proof-
 let ?l="(SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
@@ -645,7 +647,7 @@
   { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
   moreover
   { assume "~(T=UNIV)"
-    then obtain B where "T={B<..}" using T_def extreal_open_mono_set[of T] by auto
+    then obtain B where "T={B<..}" using T_def ereal_open_mono_set[of T] by auto
     hence "B<?l" using T_def by auto
     then obtain d where d_def: "0<d & B<(INF y:(S Int ball x d - {x}). f y)"
       unfolding less_SUP_iff by auto
@@ -670,14 +672,14 @@
     } hence "B <= INFI (S Int ball x d - {x}) f" apply (subst le_INFI) by auto
     also have "...<=?l" apply (subst le_SUPI) using d_def by auto
     finally have "B<=?l" by auto
-  } hence "z <= ?l" using extreal_le_extreal[of z "?l"] by auto
+  } hence "z <= ?l" using ereal_le_ereal[of z "?l"] by auto
 }
-ultimately show ?thesis unfolding extreal_Liminf_Sup_monoset eventually_within
-   apply (subst extreal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"]) by auto
+ultimately show ?thesis unfolding ereal_Liminf_Sup_monoset eventually_within
+   apply (subst ereal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"]) by auto
 qed
 
 lemma Limsup_within:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
 proof-
 let ?l="(INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
@@ -687,12 +689,12 @@
   { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
   moreover
   { assume "~(T=UNIV)" hence "~(uminus ` T = UNIV)"
-       by (metis Int_UNIV_right Int_absorb1 image_mono extreal_minus_minus_image subset_UNIV)
-    hence "uminus ` T = {Inf (uminus ` T)<..}" using T_def extreal_open_mono_set[of "uminus ` T"]
-       extreal_open_uminus[of T] by auto
+       by (metis Int_UNIV_right Int_absorb1 image_mono ereal_minus_minus_image subset_UNIV)
+    hence "uminus ` T = {Inf (uminus ` T)<..}" using T_def ereal_open_mono_set[of "uminus ` T"]
+       ereal_open_uminus[of T] by auto
     then obtain B where "T={..<B}"
-      unfolding extreal_Inf_uminus_image_eq extreal_uminus_lessThan[symmetric]
-      unfolding inj_image_eq_iff[OF extreal_inj_on_uminus] by simp
+      unfolding ereal_Inf_uminus_image_eq ereal_uminus_lessThan[symmetric]
+      unfolding inj_image_eq_iff[OF ereal_inj_on_uminus] by simp
     hence "?l<B" using T_def by auto
     then obtain d where d_def: "0<d & (SUP y:(S Int ball x d - {x}). f y)<B"
       unfolding INF_less_iff by auto
@@ -717,33 +719,33 @@
     } hence "SUPR (S Int ball x d - {x}) f <= B" apply (subst SUP_leI) by auto
     moreover have "?l<=SUPR (S Int ball x d - {x}) f" apply (subst INF_leI) using d_def by auto
     ultimately have "?l<=B" by auto
-  } hence "?l <= z" using extreal_ge_extreal[of z "?l"] by auto
+  } hence "?l <= z" using ereal_ge_ereal[of z "?l"] by auto
 }
-ultimately show ?thesis unfolding extreal_Limsup_Inf_monoset eventually_within
-   apply (subst extreal_InfI) by auto
+ultimately show ?thesis unfolding ereal_Limsup_Inf_monoset eventually_within
+   apply (subst ereal_InfI) by auto
 qed
 
 
 lemma Liminf_within_UNIV:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Liminf (at x) f = Liminf (at x within UNIV) f"
 by (metis within_UNIV)
 
 
 lemma Liminf_at:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
 using Liminf_within[of x UNIV f] Liminf_within_UNIV[of x f] by auto
 
 
 lemma Limsup_within_UNIV:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Limsup (at x) f = Limsup (at x within UNIV) f"
 by (metis within_UNIV)
 
 
 lemma Limsup_at:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
 using Limsup_within[of x UNIV f] Limsup_within_UNIV[of x f] by auto
 
@@ -755,14 +757,14 @@
 by (metis assms(1) linorder_le_less_linear n_not_Suc_n real_of_nat_le_zero_cancel_iff)
 
 lemma Liminf_within_constant:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   assumes "ALL y:S. f y = C"
   assumes "~trivial_limit (at x within S)"
   shows "Liminf (at x within S) f = C"
 by (metis Lim_within_constant assms lim_imp_Liminf)
 
 lemma Limsup_within_constant:
-  fixes f :: "'a::metric_space => extreal"
+  fixes f :: "'a::metric_space => ereal"
   assumes "ALL y:S. f y = C"
   assumes "~trivial_limit (at x within S)"
   shows "Limsup (at x within S) f = C"
@@ -805,17 +807,17 @@
 } ultimately show ?thesis by auto
 qed
 
-lemma liminf_extreal_cminus:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "c \<noteq> -\<infinity>"
+lemma liminf_ereal_cminus:
+  fixes f :: "nat \<Rightarrow> ereal" assumes "c \<noteq> -\<infinity>"
   shows "liminf (\<lambda>x. c - f x) = c - limsup f"
 proof (cases c)
   case PInf then show ?thesis by (simp add: Liminf_const)
 next
   case (real r) then show ?thesis
     unfolding liminf_SUPR_INFI limsup_INFI_SUPR
-    apply (subst INFI_extreal_cminus)
+    apply (subst INFI_ereal_cminus)
     apply auto
-    apply (subst SUPR_extreal_cminus)
+    apply (subst SUPR_ereal_cminus)
     apply auto
     done
 qed (insert `c \<noteq> -\<infinity>`, simp)
@@ -853,77 +855,77 @@
 from this show ?thesis using continuous_imp_tendsto by auto
 qed
 
-lemma continuous_at_of_extreal:
-  fixes x0 :: extreal
+lemma continuous_at_of_ereal:
+  fixes x0 :: ereal
   assumes "\<bar>x0\<bar> \<noteq> \<infinity>"
   shows "continuous (at x0) real"
 proof-
 { fix T assume T_def: "open T & real x0 : T"
-  def S == "extreal ` T"
-  hence "extreal (real x0) : S" using T_def by auto
-  hence "x0 : S" using assms extreal_real by auto
-  moreover have "open S" using open_extreal S_def T_def by auto
+  def S == "ereal ` T"
+  hence "ereal (real x0) : S" using T_def by auto
+  hence "x0 : S" using assms ereal_real by auto
+  moreover have "open S" using open_ereal S_def T_def by auto
   moreover have "ALL y:S. real y : T" using S_def T_def by auto
   ultimately have "EX S. x0 : S & open S & (ALL y:S. real y : T)" by auto
 } from this show ?thesis unfolding continuous_at_open by blast
 qed
 
 
-lemma continuous_at_iff_extreal:
+lemma continuous_at_iff_ereal:
 fixes f :: "'a::t2_space => real"
-shows "continuous (at x0) f <-> continuous (at x0) (extreal o f)"
+shows "continuous (at x0) f <-> continuous (at x0) (ereal o f)"
 proof-
-{ assume "continuous (at x0) f" hence "continuous (at x0) (extreal o f)"
-     using continuous_at_extreal continuous_at_compose[of x0 f extreal] by auto
+{ assume "continuous (at x0) f" hence "continuous (at x0) (ereal o f)"
+     using continuous_at_ereal continuous_at_compose[of x0 f ereal] by auto
 }
 moreover
-{ assume "continuous (at x0) (extreal o f)"
-  hence "continuous (at x0) (real o (extreal o f))"
-     using continuous_at_of_extreal by (intro continuous_at_compose[of x0 "extreal o f"]) auto
-  moreover have "real o (extreal o f) = f" using real_extreal_id by (simp add: o_assoc)
+{ assume "continuous (at x0) (ereal o f)"
+  hence "continuous (at x0) (real o (ereal o f))"
+     using continuous_at_of_ereal by (intro continuous_at_compose[of x0 "ereal o f"]) auto
+  moreover have "real o (ereal o f) = f" using real_ereal_id by (simp add: o_assoc)
   ultimately have "continuous (at x0) f" by auto
 } ultimately show ?thesis by auto
 qed
 
 
-lemma continuous_on_iff_extreal:
+lemma continuous_on_iff_ereal:
 fixes f :: "'a::t2_space => real"
 fixes A assumes "open A"
-shows "continuous_on A f <-> continuous_on A (extreal o f)"
-   using continuous_at_iff_extreal assms by (auto simp add: continuous_on_eq_continuous_at)
+shows "continuous_on A f <-> continuous_on A (ereal o f)"
+   using continuous_at_iff_ereal assms by (auto simp add: continuous_on_eq_continuous_at)
 
 
-lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>)}) real"
-   using continuous_at_of_extreal continuous_on_eq_continuous_at open_image_extreal by auto
+lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>::ereal)}) real"
+   using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal by auto
 
 
 lemma continuous_on_iff_real:
-  fixes f :: "'a::t2_space => extreal"
+  fixes f :: "'a::t2_space => ereal"
   assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
   shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)"
 proof-
   have "f ` A <= UNIV-{\<infinity>,(-\<infinity>)}" using assms by force
   hence *: "continuous_on (f ` A) real"
      using continuous_on_real by (simp add: continuous_on_subset)
-have **: "continuous_on ((real o f) ` A) extreal"
-   using continuous_on_extreal continuous_on_subset[of "UNIV" "extreal" "(real o f) ` A"] by blast
+have **: "continuous_on ((real o f) ` A) ereal"
+   using continuous_on_ereal continuous_on_subset[of "UNIV" "ereal" "(real o f) ` A"] by blast
 { assume "continuous_on A f" hence "continuous_on A (real o f)"
   apply (subst continuous_on_compose) using * by auto
 }
 moreover
 { assume "continuous_on A (real o f)"
-  hence "continuous_on A (extreal o (real o f))"
+  hence "continuous_on A (ereal o (real o f))"
      apply (subst continuous_on_compose) using ** by auto
   hence "continuous_on A f"
-     apply (subst continuous_on_eq[of A "extreal o (real o f)" f])
-     using assms extreal_real by auto
+     apply (subst continuous_on_eq[of A "ereal o (real o f)" f])
+     using assms ereal_real by auto
 }
 ultimately show ?thesis by auto
 qed
 
 
 lemma continuous_at_const:
-  fixes f :: "'a::t2_space => extreal"
+  fixes f :: "'a::t2_space => ereal"
   assumes "ALL x. (f x = C)"
   shows "ALL x. continuous (at x) f"
 unfolding continuous_at_open using assms t1_space by auto
@@ -977,11 +979,11 @@
 qed
 
 
-lemma mono_closed_extreal:
+lemma mono_closed_ereal:
   fixes S :: "real set"
   assumes mono: "ALL y z. y:S & y<=z --> z:S"
   assumes "closed S"
-  shows "EX a. S = {x. a <= extreal x}"
+  shows "EX a. S = {x. a <= ereal x}"
 proof-
 { assume "S = {}" hence ?thesis apply(rule_tac x=PInfty in exI) by auto }
 moreover
@@ -989,19 +991,21 @@
 moreover
 { assume "EX a. S = {a ..}"
   from this obtain a where "S={a ..}" by auto
-  hence ?thesis apply(rule_tac x="extreal a" in exI) by auto
+  hence ?thesis apply(rule_tac x="ereal a" in exI) by auto
 } ultimately show ?thesis using mono_closed_real[of S] assms by auto
 qed
 
 subsection {* Sums *}
 
-lemma setsum_extreal[simp]:
-  "(\<Sum>x\<in>A. extreal (f x)) = extreal (\<Sum>x\<in>A. f x)"
+lemma setsum_ereal[simp]:
+  "(\<Sum>x\<in>A. ereal (f x)) = ereal (\<Sum>x\<in>A. f x)"
 proof cases
   assume "finite A" then show ?thesis by induct auto
 qed simp
 
-lemma setsum_Pinfty: "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))"
+lemma setsum_Pinfty:
+  fixes f :: "'a \<Rightarrow> ereal"
+  shows "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))"
 proof safe
   assume *: "setsum f P = \<infinity>"
   show "finite P"
@@ -1023,15 +1027,16 @@
 qed
 
 lemma setsum_Inf:
+  fixes f :: "'a \<Rightarrow> ereal"
   shows "\<bar>setsum f A\<bar> = \<infinity> \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>))"
 proof
   assume *: "\<bar>setsum f A\<bar> = \<infinity>"
   have "finite A" by (rule ccontr) (insert *, auto)
   moreover have "\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>"
   proof (rule ccontr)
-    assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
+    assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto
     from bchoice[OF this] guess r ..
-    with * show False by (auto simp: setsum_extreal)
+    with * show False by (auto simp: setsum_ereal)
   qed
   ultimately show "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)" by auto
 next
@@ -1040,72 +1045,73 @@
   then show "\<bar>setsum f A\<bar> = \<infinity>"
   proof induct
     case (insert j A) then show ?case
-      by (cases rule: extreal3_cases[of "f i" "f j" "setsum f A"]) auto
+      by (cases rule: ereal3_cases[of "f i" "f j" "setsum f A"]) auto
   qed simp
 qed
 
-lemma setsum_real_of_extreal:
+lemma setsum_real_of_ereal:
+  fixes f :: "'i \<Rightarrow> ereal"
   assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
   shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
 proof -
-  have "\<forall>x\<in>S. \<exists>r. f x = extreal r"
+  have "\<forall>x\<in>S. \<exists>r. f x = ereal r"
   proof
     fix x assume "x \<in> S"
-    from assms[OF this] show "\<exists>r. f x = extreal r" by (cases "f x") auto
+    from assms[OF this] show "\<exists>r. f x = ereal r" by (cases "f x") auto
   qed
   from bchoice[OF this] guess r ..
   then show ?thesis by simp
 qed
 
-lemma setsum_extreal_0:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
+lemma setsum_ereal_0:
+  fixes f :: "'a \<Rightarrow> ereal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
   shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
 proof
   assume *: "(\<Sum>x\<in>A. f x) = 0"
   then have "(\<Sum>x\<in>A. f x) \<noteq> \<infinity>" by auto
   then have "\<forall>i\<in>A. \<bar>f i\<bar> \<noteq> \<infinity>" using assms by (force simp: setsum_Pinfty)
-  then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
+  then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto
   from bchoice[OF this] * assms show "\<forall>i\<in>A. f i = 0"
     using setsum_nonneg_eq_0_iff[of A "\<lambda>i. real (f i)"] by auto
 qed (rule setsum_0')
 
 
-lemma setsum_extreal_right_distrib:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
+lemma setsum_ereal_right_distrib:
+  fixes f :: "'a \<Rightarrow> ereal" assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
   shows "r * setsum f A = (\<Sum>n\<in>A. r * f n)"
 proof cases
   assume "finite A" then show ?thesis using assms
-    by induct (auto simp: extreal_right_distrib setsum_nonneg)
+    by induct (auto simp: ereal_right_distrib setsum_nonneg)
 qed simp
 
-lemma sums_extreal_positive:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
+lemma sums_ereal_positive:
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
 proof -
   have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)"
-    using extreal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
-  from LIMSEQ_extreal_SUPR[OF this]
+    using ereal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
+  from LIMSEQ_ereal_SUPR[OF this]
   show ?thesis unfolding sums_def by (simp add: atLeast0LessThan)
 qed
 
-lemma summable_extreal_pos:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "summable f"
-  using sums_extreal_positive[of f, OF assms] unfolding summable_def by auto
+lemma summable_ereal_pos:
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" shows "summable f"
+  using sums_ereal_positive[of f, OF assms] unfolding summable_def by auto
 
-lemma suminf_extreal_eq_SUPR:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i"
+lemma suminf_ereal_eq_SUPR:
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i"
   shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)"
-  using sums_extreal_positive[of f, OF assms, THEN sums_unique] by simp
+  using sums_ereal_positive[of f, OF assms, THEN sums_unique] by simp
 
-lemma sums_extreal:
-  "(\<lambda>x. extreal (f x)) sums extreal x \<longleftrightarrow> f sums x"
+lemma sums_ereal:
+  "(\<lambda>x. ereal (f x)) sums ereal x \<longleftrightarrow> f sums x"
   unfolding sums_def by simp
 
 lemma suminf_bound:
-  fixes f :: "nat \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> ereal"
   assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x" and pos: "\<And>n. 0 \<le> f n"
   shows "suminf f \<le> x"
-proof (rule Lim_bounded_extreal)
-  have "summable f" using pos[THEN summable_extreal_pos] .
+proof (rule Lim_bounded_ereal)
+  have "summable f" using pos[THEN summable_ereal_pos] .
   then show "(\<lambda>N. \<Sum>n<N. f n) ----> suminf f"
     by (auto dest!: summable_sums simp: sums_def atLeast0LessThan)
   show "\<forall>n\<ge>0. setsum f {..<n} \<le> x"
@@ -1113,15 +1119,15 @@
 qed
 
 lemma suminf_bound_add:
-  fixes f :: "nat \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> ereal"
   assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x" and pos: "\<And>n. 0 \<le> f n" and "y \<noteq> -\<infinity>"
   shows "suminf f + y \<le> x"
 proof (cases y)
   case (real r) then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y"
-    using assms by (simp add: extreal_le_minus)
+    using assms by (simp add: ereal_le_minus)
   then have "(\<Sum> n. f n) \<le> x - y" using pos by (rule suminf_bound)
   then show "(\<Sum> n. f n) + y \<le> x"
-    using assms real by (simp add: extreal_le_minus)
+    using assms real by (simp add: ereal_le_minus)
 qed (insert assms, auto)
 
 lemma sums_finite:
@@ -1140,22 +1146,22 @@
   shows "suminf f = (\<Sum>N<n. f N)"
   using sums_finite[OF assms, THEN sums_unique] by simp
 
-lemma suminf_extreal_0[simp]: "(\<Sum>i. 0) = (0::'a::{comm_monoid_add,t2_space})"
+lemma suminf_ereal_0[simp]: "(\<Sum>i. 0) = (0::'a::{comm_monoid_add,t2_space})"
   using suminf_finite[of 0 "\<lambda>x. 0"] by simp
 
 lemma suminf_upper:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>n. 0 \<le> f n"
   shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)"
-  unfolding suminf_extreal_eq_SUPR[OF assms] SUPR_def
+  unfolding suminf_ereal_eq_SUPR[OF assms] SUPR_def
   by (auto intro: complete_lattice_class.Sup_upper image_eqI)
 
 lemma suminf_0_le:
-  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
+  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>n. 0 \<le> f n"
   shows "0 \<le> (\<Sum>n. f n)"
   using suminf_upper[of f 0, OF assms] by simp
 
 lemma suminf_le_pos:
-  fixes f g :: "nat \<Rightarrow> extreal"
+  fixes f g :: "nat \<Rightarrow> ereal"
   assumes "\<And>N. f N \<le> g N" "\<And>N. 0 \<le> f N"
   shows "suminf f \<le> suminf g"
 proof (safe intro!: suminf_bound)
@@ -1165,27 +1171,28 @@
   finally show "setsum f {..<n} \<le> suminf g" .
 qed (rule assms(2))
 
-lemma suminf_half_series_extreal: "(\<Sum>n. (1/2 :: extreal)^Suc n) = 1"
-  using sums_extreal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]
-  by (simp add: one_extreal_def)
+lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal)^Suc n) = 1"
+  using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]
+  by (simp add: one_ereal_def)
 
-lemma suminf_add_extreal:
-  fixes f g :: "nat \<Rightarrow> extreal"
+lemma suminf_add_ereal:
+  fixes f g :: "nat \<Rightarrow> ereal"
   assumes "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
   shows "(\<Sum>i. f i + g i) = suminf f + suminf g"
-  apply (subst (1 2 3) suminf_extreal_eq_SUPR)
+  apply (subst (1 2 3) suminf_ereal_eq_SUPR)
   unfolding setsum_addf
-  by (intro assms extreal_add_nonneg_nonneg SUPR_extreal_add_pos incseq_setsumI setsum_nonneg ballI)+
+  by (intro assms ereal_add_nonneg_nonneg SUPR_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+
 
-lemma suminf_cmult_extreal:
-  fixes f g :: "nat \<Rightarrow> extreal"
+lemma suminf_cmult_ereal:
+  fixes f g :: "nat \<Rightarrow> ereal"
   assumes "\<And>i. 0 \<le> f i" "0 \<le> a"
   shows "(\<Sum>i. a * f i) = a * suminf f"
-  by (auto simp: setsum_extreal_right_distrib[symmetric] assms
-                 extreal_zero_le_0_iff setsum_nonneg suminf_extreal_eq_SUPR
-           intro!: SUPR_extreal_cmult )
+  by (auto simp: setsum_ereal_right_distrib[symmetric] assms
+                 ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUPR
+           intro!: SUPR_ereal_cmult )
 
 lemma suminf_PInfty:
+  fixes f :: "nat \<Rightarrow> ereal"
   assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
   shows "f i \<noteq> \<infinity>"
 proof -
@@ -1197,43 +1204,43 @@
 
 lemma suminf_PInfty_fun:
   assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
-  shows "\<exists>f'. f = (\<lambda>x. extreal (f' x))"
+  shows "\<exists>f'. f = (\<lambda>x. ereal (f' x))"
 proof -
-  have "\<forall>i. \<exists>r. f i = extreal r"
+  have "\<forall>i. \<exists>r. f i = ereal r"
   proof
-    fix i show "\<exists>r. f i = extreal r"
+    fix i show "\<exists>r. f i = ereal r"
       using suminf_PInfty[OF assms] assms(1)[of i] by (cases "f i") auto
   qed
   from choice[OF this] show ?thesis by auto
 qed
 
-lemma summable_extreal:
-  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
+lemma summable_ereal:
+  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"
   shows "summable f"
 proof -
-  have "0 \<le> (\<Sum>i. extreal (f i))"
+  have "0 \<le> (\<Sum>i. ereal (f i))"
     using assms by (intro suminf_0_le) auto
-  with assms obtain r where r: "(\<Sum>i. extreal (f i)) = extreal r"
-    by (cases "\<Sum>i. extreal (f i)") auto
-  from summable_extreal_pos[of "\<lambda>x. extreal (f x)"]
-  have "summable (\<lambda>x. extreal (f x))" using assms by auto
+  with assms obtain r where r: "(\<Sum>i. ereal (f i)) = ereal r"
+    by (cases "\<Sum>i. ereal (f i)") auto
+  from summable_ereal_pos[of "\<lambda>x. ereal (f x)"]
+  have "summable (\<lambda>x. ereal (f x))" using assms by auto
   from summable_sums[OF this]
-  have "(\<lambda>x. extreal (f x)) sums (\<Sum>x. extreal (f x))" by auto
+  have "(\<lambda>x. ereal (f x)) sums (\<Sum>x. ereal (f x))" by auto
   then show "summable f"
-    unfolding r sums_extreal summable_def ..
+    unfolding r sums_ereal summable_def ..
 qed
 
-lemma suminf_extreal:
-  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
-  shows "(\<Sum>i. extreal (f i)) = extreal (suminf f)"
+lemma suminf_ereal:
+  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"
+  shows "(\<Sum>i. ereal (f i)) = ereal (suminf f)"
 proof (rule sums_unique[symmetric])
-  from summable_extreal[OF assms]
-  show "(\<lambda>x. extreal (f x)) sums (extreal (suminf f))"
-    unfolding sums_extreal using assms by (intro summable_sums summable_extreal)
+  from summable_ereal[OF assms]
+  show "(\<lambda>x. ereal (f x)) sums (ereal (suminf f))"
+    unfolding sums_ereal using assms by (intro summable_sums summable_ereal)
 qed
 
-lemma suminf_extreal_minus:
-  fixes f g :: "nat \<Rightarrow> extreal"
+lemma suminf_ereal_minus:
+  fixes f g :: "nat \<Rightarrow> ereal"
   assumes ord: "\<And>i. g i \<le> f i" "\<And>i. 0 \<le> g i" and fin: "suminf f \<noteq> \<infinity>" "suminf g \<noteq> \<infinity>"
   shows "(\<Sum>i. f i - g i) = suminf f - suminf g"
 proof -
@@ -1241,50 +1248,51 @@
   moreover
   from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] guess f' .. note this[simp]
   from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] guess g' .. note this[simp]
-  { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: extreal_le_minus_iff) }
+  { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: ereal_le_minus_iff) }
   moreover
   have "suminf (\<lambda>i. f i - g i) \<le> suminf f"
     using assms by (auto intro!: suminf_le_pos simp: field_simps)
   then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>" using fin by auto
   ultimately show ?thesis using assms `\<And>i. 0 \<le> f i`
     apply simp
-    by (subst (1 2 3) suminf_extreal)
-       (auto intro!: suminf_diff[symmetric] summable_extreal)
+    by (subst (1 2 3) suminf_ereal)
+       (auto intro!: suminf_diff[symmetric] summable_ereal)
 qed
 
-lemma suminf_extreal_PInf[simp]:
-  "(\<Sum>x. \<infinity>) = \<infinity>"
+lemma suminf_ereal_PInf[simp]:
+  "(\<Sum>x. \<infinity>::ereal) = \<infinity>"
 proof -
-  have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>)" by (rule suminf_upper) auto
+  have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>::ereal)" by (rule suminf_upper) auto
   then show ?thesis by simp
 qed
 
-lemma summable_real_of_extreal:
+lemma summable_real_of_ereal:
+  fixes f :: "nat \<Rightarrow> ereal"
   assumes f: "\<And>i. 0 \<le> f i" and fin: "(\<Sum>i. f i) \<noteq> \<infinity>"
   shows "summable (\<lambda>i. real (f i))"
 proof (rule summable_def[THEN iffD2])
   have "0 \<le> (\<Sum>i. f i)" using assms by (auto intro: suminf_0_le)
-  with fin obtain r where r: "extreal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto
+  with fin obtain r where r: "ereal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto
   { fix i have "f i \<noteq> \<infinity>" using f by (intro suminf_PInfty[OF _ fin]) auto
     then have "\<bar>f i\<bar> \<noteq> \<infinity>" using f[of i] by auto }
   note fin = this
-  have "(\<lambda>i. extreal (real (f i))) sums (\<Sum>i. extreal (real (f i)))"
-    using f by (auto intro!: summable_extreal_pos summable_sums simp: extreal_le_real_iff zero_extreal_def)
-  also have "\<dots> = extreal r" using fin r by (auto simp: extreal_real)
-  finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_extreal)
+  have "(\<lambda>i. ereal (real (f i))) sums (\<Sum>i. ereal (real (f i)))"
+    using f by (auto intro!: summable_ereal_pos summable_sums simp: ereal_le_real_iff zero_ereal_def)
+  also have "\<dots> = ereal r" using fin r by (auto simp: ereal_real)
+  finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_ereal)
 qed
 
 lemma suminf_SUP_eq:
-  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ereal"
   assumes "\<And>i. incseq (\<lambda>n. f n i)" "\<And>n i. 0 \<le> f n i"
   shows "(\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)"
 proof -
   { fix n :: nat
     have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
-      using assms by (auto intro!: SUPR_extreal_setsum[symmetric]) }
+      using assms by (auto intro!: SUPR_ereal_setsum[symmetric]) }
   note * = this
   show ?thesis using assms
-    apply (subst (1 2) suminf_extreal_eq_SUPR)
+    apply (subst (1 2) suminf_ereal_eq_SUPR)
     unfolding *
     apply (auto intro!: le_SUPI2)
     apply (subst SUP_commute) ..
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -291,7 +291,7 @@
         (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
       by (auto intro!: M2.measure_compl simp: vimage_Diff)
     with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
-      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff)
+      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_ereal_diff)
   next
     fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
     moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
@@ -401,7 +401,7 @@
   apply (simp add: pair_measure_def pair_measure_generator_def)
 proof (rule M1.positive_integral_cong)
   fix x assume "x \<in> space M1"
-  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)"
+  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: ereal)"
     unfolding indicator_def by auto
   show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
     unfolding *
@@ -656,7 +656,7 @@
   show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
     and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
     by (auto simp del: vimage_Int cong: measurable_cong
-             intro!: M1.borel_measurable_extreal_setsum setsum_cong
+             intro!: M1.borel_measurable_ereal_setsum setsum_cong
              simp add: M1.positive_integral_setsum simple_integral_def
                        M1.positive_integral_cmult
                        M1.positive_integral_cong[OF eq]
@@ -760,7 +760,7 @@
     show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
       by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
     show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N)
+      by (intro M1.borel_measurable_ereal_neq_const measure_cut_measurable_fst N)
     { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
       have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
       proof (rule M2.AE_I)
@@ -822,45 +822,45 @@
   shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
     and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
 proof -
-  let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)"
+  let "?pf x" = "ereal (f x)" and "?nf x" = "ereal (- f x)"
   have
     borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
     int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
     using assms by auto
-  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
-     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
     using borel[THEN positive_integral_fst_measurable(1)] int
     unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
   with borel[THEN positive_integral_fst_measurable(1)]
-  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
-    "AE x in M1. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+    "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
     by (auto intro!: M1.positive_integral_PInf_AE )
-  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
-    "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+    "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
     by (auto simp: M2.positive_integral_positive)
   from AE_pos show ?AE using assms
     by (simp add: measurable_pair_image_snd integrable_def)
-  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
       using M2.positive_integral_positive
-      by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder)
-    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
+      by (intro M1.positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
+    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
   note this[simp]
-  { fix f assume borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable P"
-      and int: "integral\<^isup>P P (\<lambda>x. extreal (f x)) \<noteq> \<infinity>"
-      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
-    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
+  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable P"
+      and int: "integral\<^isup>P P (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
+      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
+    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
     proof (intro integrable_def[THEN iffD2] conjI)
       show "?f \<in> borel_measurable M1"
-        using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable)
-      have "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y))  \<partial>M2) \<partial>M1)"
+        using borel by (auto intro!: M1.borel_measurable_real_of_ereal positive_integral_fst_measurable)
+      have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y))  \<partial>M2) \<partial>M1)"
         using AE M2.positive_integral_positive
-        by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real)
-      then show "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) \<noteq> \<infinity>"
+        by (auto intro!: M1.positive_integral_cong_AE simp: ereal_real)
+      then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
         using positive_integral_fst_measurable[OF borel] int by simp
-      have "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+      have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
         by (intro M1.positive_integral_cong_pos)
-           (simp add: M2.positive_integral_positive real_of_extreal_pos)
-      then show "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
+           (simp add: M2.positive_integral_positive real_of_ereal_pos)
+      then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
     qed }
   with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
   show ?INT
--- a/src/HOL/Probability/Borel_Space.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Borel_Space.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -112,7 +112,7 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_restricted:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
+  fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M"
   shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
     (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
     (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
@@ -123,7 +123,7 @@
   show ?thesis unfolding *
     unfolding in_borel_measurable_borel
   proof (simp, safe)
-    fix S :: "extreal set" assume "S \<in> sets borel"
+    fix S :: "ereal set" assume "S \<in> sets borel"
       "\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
     then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
     then have f: "?f -` S \<inter> A \<in> sets M"
@@ -142,7 +142,7 @@
       then show ?thesis using f by auto
     qed
   next
-    fix S :: "extreal set" assume "S \<in> sets borel"
+    fix S :: "ereal set" assume "S \<in> sets borel"
       "\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M"
     then have f: "?f -` S \<inter> space M \<in> sets M" by auto
     then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
@@ -1095,70 +1095,71 @@
 
 subsection "Borel space on the extended reals"
 
-lemma borel_measurable_extreal_borel:
-  "extreal \<in> borel_measurable borel"
-  unfolding borel_def[where 'a=extreal]
+lemma borel_measurable_ereal_borel:
+  "ereal \<in> borel_measurable borel"
+  unfolding borel_def[where 'a=ereal]
 proof (rule borel.measurable_sigma)
-  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
+  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
   then have "open X" by (auto simp: mem_def)
-  then have "open (extreal -` X \<inter> space borel)"
-    by (simp add: open_extreal_vimage)
-  then show "extreal -` X \<inter> space borel \<in> sets borel" by auto
+  then have "open (ereal -` X \<inter> space borel)"
+    by (simp add: open_ereal_vimage)
+  then show "ereal -` X \<inter> space borel \<in> sets borel" by auto
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_extreal[simp, intro]:
-  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
-  using measurable_comp[OF f borel_measurable_extreal_borel] unfolding comp_def .
+lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]:
+  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
+  using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def .
 
-lemma borel_measurable_real_of_extreal_borel:
-  "(real :: extreal \<Rightarrow> real) \<in> borel_measurable borel"
+lemma borel_measurable_real_of_ereal_borel:
+  "(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel"
   unfolding borel_def[where 'a=real]
 proof (rule borel.measurable_sigma)
   fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
   then have "open B" by (auto simp: mem_def)
-  have *: "extreal -` real -` (B - {0}) = B - {0}" by auto
-  have open_real: "open (real -` (B - {0}) :: extreal set)"
-    unfolding open_extreal_def * using `open B` by auto
-  show "(real -` B \<inter> space borel :: extreal set) \<in> sets borel"
+  have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
+  have open_real: "open (real -` (B - {0}) :: ereal set)"
+    unfolding open_ereal_def * using `open B` by auto
+  show "(real -` B \<inter> space borel :: ereal set) \<in> sets borel"
   proof cases
     assume "0 \<in> B"
-    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0}"
-      by (auto simp add: real_of_extreal_eq_0)
-    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0::ereal}"
+      by (auto simp add: real_of_ereal_eq_0)
+    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
       using open_real by auto
   next
     assume "0 \<notin> B"
-    then have *: "(real -` B :: extreal set) = real -` (B - {0})"
-      by (auto simp add: real_of_extreal_eq_0)
-    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
+    then have *: "(real -` B :: ereal set) = real -` (B - {0})"
+      by (auto simp add: real_of_ereal_eq_0)
+    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
       using open_real by auto
   qed
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_real_of_extreal[simp, intro]:
-  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: extreal)) \<in> borel_measurable M"
-  using measurable_comp[OF f borel_measurable_real_of_extreal_borel] unfolding comp_def .
+lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]:
+  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M"
+  using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def .
 
-lemma (in sigma_algebra) borel_measurable_extreal_iff:
-  shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_ereal_iff:
+  shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
 proof
-  assume "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
-  from borel_measurable_real_of_extreal[OF this]
+  assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
+  from borel_measurable_real_of_ereal[OF this]
   show "f \<in> borel_measurable M" by auto
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_extreal_iff_real:
-  "f \<in> borel_measurable M \<longleftrightarrow>
+lemma (in sigma_algebra) borel_measurable_ereal_iff_real:
+  fixes f :: "'a \<Rightarrow> ereal"
+  shows "f \<in> borel_measurable M \<longleftrightarrow>
     ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
 proof safe
   assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
   have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
   with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
-  let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else extreal (real (f x))"
+  let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real (f x))"
   have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
-  also have "?f = f" by (auto simp: fun_eq_iff extreal_real)
+  also have "?f = f" by (auto simp: fun_eq_iff ereal_real)
   finally show "f \<in> borel_measurable M" .
-qed (auto intro: measurable_sets borel_measurable_real_of_extreal)
+qed (auto intro: measurable_sets borel_measurable_real_of_ereal)
 
 lemma (in sigma_algebra) less_eq_ge_measurable:
   fixes f :: "'a \<Rightarrow> 'c::linorder"
@@ -1186,95 +1187,96 @@
   ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_uminus_borel_extreal:
-  "(uminus :: extreal \<Rightarrow> extreal) \<in> borel_measurable borel"
+lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal:
+  "(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel"
 proof (subst borel_def, rule borel.measurable_sigma)
-  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
+  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
   then have "open X" by (simp add: mem_def)
   have "uminus -` X = uminus ` X" by (force simp: image_iff)
-  then have "open (uminus -` X)" using `open X` extreal_open_uminus by auto
+  then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto
   then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_uminus_extreal[intro]:
+lemma (in sigma_algebra) borel_measurable_uminus_ereal[intro]:
   assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M"
-  using measurable_comp[OF assms borel_measurable_uminus_borel_extreal] by (simp add: comp_def)
+  shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
+  using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def)
 
-lemma (in sigma_algebra) borel_measurable_uminus_eq_extreal[simp]:
-  "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
+lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]:
+  "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
 proof
-  assume ?l from borel_measurable_uminus_extreal[OF this] show ?r by simp
+  assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_eq_atMost_extreal:
-  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
+lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal:
+  fixes f :: "'a \<Rightarrow> ereal"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
 proof (intro iffI allI)
   assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
   show "f \<in> borel_measurable M"
-    unfolding borel_measurable_extreal_iff_real borel_measurable_iff_le
+    unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le
   proof (intro conjI allI)
     fix a :: real
-    { fix x :: extreal assume *: "\<forall>i::nat. real i < x"
+    { fix x :: ereal assume *: "\<forall>i::nat. real i < x"
       have "x = \<infinity>"
-      proof (rule extreal_top)
+      proof (rule ereal_top)
         fix B from real_arch_lt[of B] guess n ..
-        then have "extreal B < real n" by auto
+        then have "ereal B < real n" by auto
         with * show "B \<le> x" by (metis less_trans less_imp_le)
       qed }
     then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
       by (auto simp: not_le)
     then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
     moreover
-    have "{-\<infinity>} = {..-\<infinity>}" by auto
+    have "{-\<infinity>::ereal} = {..-\<infinity>}" by auto
     then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
-    moreover have "{x\<in>space M. f x \<le> extreal a} \<in> sets M"
-      using pos[of "extreal a"] by (simp add: vimage_def Int_def conj_commute)
+    moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
+      using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute)
     moreover have "{w \<in> space M. real (f w) \<le> a} =
-      (if a < 0 then {w \<in> space M. f w \<le> extreal a} - f -` {-\<infinity>} \<inter> space M
-      else {w \<in> space M. f w \<le> extreal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
+      (if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
+      else {w \<in> space M. f w \<le> ereal a} \<union> (f -` {\<infinity>} \<inter> space M) \<union> (f -` {-\<infinity>} \<inter> space M))" (is "?l = ?r")
       proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed
     ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
   qed
 qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_eq_atLeast_extreal:
-  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
+lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal:
+  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
 proof
   assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
   moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
-    by (auto simp: extreal_uminus_le_reorder)
+    by (auto simp: ereal_uminus_le_reorder)
   ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
-    unfolding borel_measurable_eq_atMost_extreal by auto
+    unfolding borel_measurable_eq_atMost_ereal by auto
   then show "f \<in> borel_measurable M" by simp
 qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_extreal_iff_less:
-  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
-  unfolding borel_measurable_eq_atLeast_extreal greater_eq_le_measurable ..
+lemma (in sigma_algebra) borel_measurable_ereal_iff_less:
+  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
+  unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_extreal_iff_ge:
-  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
-  unfolding borel_measurable_eq_atMost_extreal less_eq_ge_measurable ..
+lemma (in sigma_algebra) borel_measurable_ereal_iff_ge:
+  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
+  unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_extreal_eq_const:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_ereal_eq_const:
+  fixes f :: "'a \<Rightarrow> ereal" 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_extreal_neq_const:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_ereal_neq_const:
+  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
   shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 proof -
   have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
   then show ?thesis using assms by (auto intro!: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_le[intro,simp]:
-  fixes f g :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]:
+  fixes f g :: "'a \<Rightarrow> ereal"
   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"
@@ -1283,13 +1285,13 @@
     {x \<in> space M. real (f x) \<le> real (g x)} - (f -` {\<infinity>, -\<infinity>} \<inter> space M \<union> g -` {\<infinity>, -\<infinity>} \<inter> space M) \<union>
     f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
   proof (intro set_eqI)
-    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: extreal2_cases[of "f x" "g x"]) auto
+    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: ereal2_cases[of "f x" "g x"]) auto
   qed
   with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_less[intro,simp]:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]:
+  fixes f :: "'a \<Rightarrow> ereal"
   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"
@@ -1298,8 +1300,8 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_eq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> ereal"
   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"
@@ -1308,8 +1310,8 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_neq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> ereal"
   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"
@@ -1323,23 +1325,23 @@
   "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
   by auto
 
-lemma (in sigma_algebra) borel_measurable_extreal_add[intro, simp]:
-  fixes f :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]:
+  fixes f :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
 proof -
   { fix x assume "x \<in> space M" then have "f x + g x =
       (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
         else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>
-        else extreal (real (f x) + real (g x)))"
-      by (cases rule: extreal2_cases[of "f x" "g x"]) auto }
+        else ereal (real (f x) + real (g x)))"
+      by (cases rule: ereal2_cases[of "f x" "g x"]) auto }
   with assms show ?thesis
     by (auto cong: measurable_cong simp: split_sets
              intro!: Un measurable_If measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_setsum[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   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
@@ -1348,25 +1350,25 @@
     by induct auto
 qed (simp add: borel_measurable_const)
 
-lemma (in sigma_algebra) borel_measurable_extreal_abs[intro, simp]:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]:
+  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
 proof -
   { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
   then show ?thesis using assms by (auto intro!: measurable_If)
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_times[intro, simp]:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]:
+  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
 proof -
-  { fix f g :: "'a \<Rightarrow> extreal"
+  { fix f g :: "'a \<Rightarrow> ereal"
     assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
       and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x"
     { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
         else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
-        else extreal (real (f x) * real (g x)))"
-      apply (cases rule: extreal2_cases[of "f x" "g x"])
+        else ereal (real (f x) * real (g x)))"
+      apply (cases rule: ereal2_cases[of "f x" "g x"])
       using pos[of x] by auto }
     with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M"
       by (auto cong: measurable_cong simp: split_sets
@@ -1376,12 +1378,12 @@
     (\<lambda>x. if 0 \<le> f x \<and> 0 \<le> g x \<or> f x < 0 \<and> g x < 0 then \<bar>f x\<bar> * \<bar>g x\<bar> else - (\<bar>f x\<bar> * \<bar>g x\<bar>))"
     by (auto simp: fun_eq_iff)
   show ?thesis using assms unfolding *
-    by (intro measurable_If pos_times borel_measurable_uminus_extreal)
+    by (intro measurable_If pos_times borel_measurable_uminus_ereal)
        (auto simp: split_sets intro!: Int)
 qed
 
-lemma (in sigma_algebra) borel_measurable_extreal_setprod[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
 proof cases
@@ -1389,25 +1391,25 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma (in sigma_algebra) borel_measurable_extreal_min[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> ereal"
   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_extreal_max[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> ereal"
   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> extreal"
+  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
-  unfolding borel_measurable_extreal_iff_ge
+  unfolding borel_measurable_ereal_iff_ge
 proof
   fix a
   have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
@@ -1417,10 +1419,10 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
-  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
-  unfolding borel_measurable_extreal_iff_less
+  unfolding borel_measurable_ereal_iff_less
 proof
   fix a
   have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
@@ -1430,30 +1432,30 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding liminf_SUPR_INFI using assms by auto
 
 lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding limsup_INFI_SUPR using assms by auto
 
-lemma (in sigma_algebra) borel_measurable_extreal_diff[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> extreal"
+lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  unfolding minus_extreal_def using assms by auto
+  unfolding minus_ereal_def using assms by auto
 
 lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
   shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
   apply (subst measurable_cong)
-  apply (subst suminf_extreal_eq_SUPR)
+  apply (subst suminf_ereal_eq_SUPR)
   apply (rule pos)
   using assms by auto
 
@@ -1465,11 +1467,11 @@
   and u: "\<And>i. u i \<in> borel_measurable M"
   shows "u' \<in> borel_measurable M"
 proof -
-  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. extreal (u n x)) = extreal (u' x)"
-    using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_extreal)
-  moreover from u have "(\<lambda>x. liminf (\<lambda>n. extreal (u n x))) \<in> borel_measurable M"
+  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
+    using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_ereal)
+  moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
     by auto
-  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_extreal_iff)
+  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
 qed
 
 end
--- a/src/HOL/Probability/Caratheodory.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -19,8 +19,8 @@
   Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
 *}
 
-lemma suminf_extreal_2dimen:
-  fixes f:: "nat \<times> nat \<Rightarrow> extreal"
+lemma suminf_ereal_2dimen:
+  fixes f:: "nat \<times> nat \<Rightarrow> ereal"
   assumes pos: "\<And>p. 0 \<le> f p"
   assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
   shows "(\<Sum>i. f (prod_decode i)) = suminf g"
@@ -47,21 +47,21 @@
   ultimately
   show ?thesis unfolding g_def using pos
     by (auto intro!: SUPR_eq  simp: setsum_cartesian_product reindex le_SUPI2
-                     setsum_nonneg suminf_extreal_eq_SUPR SUPR_pair
-                     SUPR_extreal_setsum[symmetric] incseq_setsumI setsum_nonneg)
+                     setsum_nonneg suminf_ereal_eq_SUPR SUPR_pair
+                     SUPR_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
 qed
 
 subsection {* Measure Spaces *}
 
 record 'a measure_space = "'a algebra" +
-  measure :: "'a set \<Rightarrow> extreal"
+  measure :: "'a set \<Rightarrow> ereal"
 
-definition positive where "positive M f \<longleftrightarrow> f {} = (0::extreal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
+definition positive where "positive M f \<longleftrightarrow> f {} = (0::ereal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
 
 definition additive where "additive M f \<longleftrightarrow>
   (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
 
-definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> extreal) \<Rightarrow> bool" where
+definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
   "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
     (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
 
@@ -168,7 +168,7 @@
   by (simp add: lambda_system_def)
 
 lemma (in algebra) lambda_system_Compl:
-  fixes f:: "'a set \<Rightarrow> extreal"
+  fixes f:: "'a set \<Rightarrow> ereal"
   assumes x: "x \<in> lambda_system M f"
   shows "space M - x \<in> lambda_system M f"
 proof -
@@ -181,7 +181,7 @@
 qed
 
 lemma (in algebra) lambda_system_Int:
-  fixes f:: "'a set \<Rightarrow> extreal"
+  fixes f:: "'a set \<Rightarrow> ereal"
   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 -
@@ -215,7 +215,7 @@
 qed
 
 lemma (in algebra) lambda_system_Un:
-  fixes f:: "'a set \<Rightarrow> extreal"
+  fixes f:: "'a set \<Rightarrow> ereal"
   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 -
@@ -321,7 +321,7 @@
 qed
 
 lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> extreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
   assumes f: "positive M f" and ad: "additive M f"
       and inc: "increasing M f"
       and A: "range A \<subseteq> sets M"
@@ -346,7 +346,7 @@
   by (simp add: positive_def lambda_system_def)
 
 lemma (in algebra) lambda_system_strong_sum:
-  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> extreal"
+  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
   assumes f: "positive M f" and a: "a \<in> sets M"
       and A: "range A \<subseteq> lambda_system M f"
       and disj: "disjoint_family A"
@@ -537,7 +537,7 @@
   assumes posf: "positive M f" and ca: "countably_additive M f"
       and s: "s \<in> sets M"
   shows "Inf (measure_set M f s) = f s"
-  unfolding Inf_extreal_def
+  unfolding Inf_ereal_def
 proof (safe intro!: Greatest_equality)
   fix z
   assume z: "z \<in> measure_set M f s"
@@ -648,7 +648,7 @@
 qed
 
 lemma (in ring_of_sets) inf_measure_close:
-  fixes e :: extreal
+  fixes e :: ereal
   assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
   shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
                (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
@@ -656,7 +656,7 @@
   from `Inf (measure_set M f s) \<noteq> \<infinity>` have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
     using inf_measure_pos[OF posf, of s] by auto
   obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
-    using Inf_extreal_close[OF fin e] by auto
+    using Inf_ereal_close[OF fin e] by auto
   thus ?thesis
     by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
 qed
@@ -672,11 +672,11 @@
      and disj: "disjoint_family A"
      and sb: "(\<Union>i. A i) \<subseteq> space M"
 
-  { fix e :: extreal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
+  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
     hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
         A n \<subseteq> (\<Union>i. BB n i) \<and> (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
       apply (safe intro!: choice inf_measure_close [of f, OF posf])
-      using e sb by (auto simp: extreal_zero_less_0_iff one_extreal_def)
+      using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def)
     then obtain BB
       where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
       and disjBB: "\<And>n. disjoint_family (BB n)"
@@ -686,15 +686,15 @@
     have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n)) + e"
     proof -
       have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
-        using suminf_half_series_extreal e
-        by (simp add: extreal_zero_le_0_iff zero_le_divide_extreal suminf_cmult_extreal)
+        using suminf_half_series_ereal e
+        by (simp add: ereal_zero_le_0_iff zero_le_divide_ereal suminf_cmult_ereal)
       have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto
       then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le)
       then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n) + e*(1/2) ^ Suc n)"
         by (rule suminf_le_pos[OF BBle])
       also have "... = (\<Sum>n. ?outer (A n)) + e"
         using sum_eq_1 inf_measure_pos[OF posf] e
-        by (subst suminf_add_extreal) (auto simp add: extreal_zero_le_0_iff)
+        by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff)
       finally show ?thesis .
     qed
     def C \<equiv> "(split BB) o prod_decode"
@@ -716,7 +716,7 @@
       by (rule ext)  (auto simp add: C_def)
     moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle
       using BB posf[unfolded positive_def]
-      by (force intro!: suminf_extreal_2dimen simp: o_def)
+      by (force intro!: suminf_ereal_2dimen simp: o_def)
     ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def)
     have "?outer (\<Union>i. A i) \<le> (\<Sum>n. \<Sum>i. f (BB n i))"
       apply (rule inf_measure_le [OF posf(1) inc], auto)
@@ -732,7 +732,7 @@
   proof cases
     assume "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
     with for_finite_Inf show ?thesis
-      by (intro extreal_le_epsilon) auto
+      by (intro ereal_le_epsilon) auto
   next
     assume "\<not> (\<forall>i. ?outer (A i) \<noteq> \<infinity>)"
     then have "\<exists>i. ?outer (A i) = \<infinity>"
@@ -771,7 +771,7 @@
   next
     assume fin: "Inf (measure_set M f s) \<noteq> \<infinity>"
     then have "measure_set M f s \<noteq> {}"
-      by (auto simp: top_extreal_def)
+      by (auto simp: top_ereal_def)
     show ?thesis
     proof (rule complete_lattice_class.Inf_greatest)
       fix r assume "r \<in> measure_set M f s"
@@ -793,7 +793,7 @@
       ultimately have "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le>
           (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
       also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
-        using A(2) x posf by (subst suminf_add_extreal) (auto simp: positive_def)
+        using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def)
       also have "\<dots> = (\<Sum>i. f (A i))"
         using A x
         by (subst add[THEN additiveD, symmetric])
@@ -830,7 +830,7 @@
 
 theorem (in ring_of_sets) caratheodory:
   assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
             measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
 proof -
   have inc: "increasing M f"
@@ -873,7 +873,7 @@
     by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
   moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
     using f(1)[unfolded positive_def] dA
-    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_extreal_pos)
+    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
   from LIMSEQ_Suc[OF this]
   have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
     unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
@@ -936,13 +936,13 @@
     show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> sets M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
       using A by auto
   qed
-  from INF_Lim_extreal[OF decseq_f this]
+  from INF_Lim_ereal[OF decseq_f this]
   have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
   moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
     by auto
   ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
     using A(4) f_fin f_Int_fin
-    by (subst INFI_extreal_add) (auto simp: decseq_f)
+    by (subst INFI_ereal_add) (auto simp: decseq_f)
   moreover {
     fix n
     have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
@@ -952,7 +952,7 @@
     finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
   ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
     by simp
-  with LIMSEQ_extreal_INFI[OF decseq_fA]
+  with LIMSEQ_ereal_INFI[OF decseq_fA]
   show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
 qed
 
@@ -965,9 +965,9 @@
   assumes A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
   shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
 proof -
-  have "\<forall>A\<in>sets M. \<exists>x. f A = extreal x"
+  have "\<forall>A\<in>sets M. \<exists>x. f A = ereal x"
   proof
-    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = extreal x"
+    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = ereal x"
       unfolding positive_def by (cases "f A") auto
   qed
   from bchoice[OF this] guess f' .. note f' = this[rule_format]
@@ -981,10 +981,10 @@
       by auto
     finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
       using A by (subst (asm) (1 2 3) f') auto
-    then have "f ((\<Union>i. A i) - A i) = extreal (f' (\<Union>i. A i) - f' (A i))"
+    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
       using A f' by auto }
   ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
-    by (simp add: zero_extreal_def)
+    by (simp add: zero_ereal_def)
   then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
     by (rule LIMSEQ_diff_approach_zero2[OF LIMSEQ_const])
   then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
@@ -1002,7 +1002,7 @@
 lemma (in ring_of_sets) caratheodory_empty_continuous:
   assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> sets M \<Longrightarrow> f A \<noteq> \<infinity>"
   assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
             measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
 proof (intro caratheodory empty_continuous_imp_countably_additive f)
   show "\<forall>A\<in>sets M. f A \<noteq> \<infinity>" using fin by auto
--- a/src/HOL/Probability/Complete_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Complete_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -253,7 +253,7 @@
 qed
 
 lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
-  fixes g :: "'a \<Rightarrow> extreal"
+  fixes g :: "'a \<Rightarrow> ereal"
   assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
   shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
 proof -
@@ -279,7 +279,7 @@
 qed
 
 lemma (in completeable_measure_space) completion_ex_borel_measurable:
-  fixes g :: "'a \<Rightarrow> extreal"
+  fixes g :: "'a \<Rightarrow> ereal"
   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/Conditional_Probability.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Conditional_Probability.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -15,7 +15,7 @@
     \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
 
 lemma (in prob_space) conditional_expectation_exists:
-  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
   assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
   shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
@@ -52,7 +52,7 @@
 qed
 
 lemma (in prob_space)
-  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
+  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
   assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
   shows borel_measurable_conditional_expectation:
@@ -71,7 +71,7 @@
 qed
 
 lemma (in sigma_algebra) factorize_measurable_function_pos:
-  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
+  fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
   assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
   assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
   shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
@@ -98,7 +98,7 @@
       show "simple_function M' ?g" using B by auto
 
       fix x assume "x \<in> space M"
-      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::extreal)"
+      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::ereal)"
         unfolding indicator_def using B by auto
       then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
         by (subst va.simple_function_indicator_representation) auto
@@ -119,7 +119,7 @@
 qed
 
 lemma (in sigma_algebra) factorize_measurable_function:
-  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
+  fixes Z :: "'a \<Rightarrow> ereal" 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))"
@@ -129,7 +129,7 @@
   from M'.sigma_algebra_vimage[OF this]
   interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 
-  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
+  { fix g :: "'c \<Rightarrow> ereal" 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)
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -538,13 +538,13 @@
 next
   assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
   then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
-    by (auto simp: setprod_extreal_0 intro!: bexI)
+    by (auto simp: setprod_ereal_0 intro!: bexI)
   moreover
   have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
     by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
        (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
   then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
-    by (auto simp: setprod_extreal_0 intro!: bexI)
+    by (auto simp: setprod_ereal_0 intro!: bexI)
   ultimately show ?thesis
     unfolding product_algebra_generator_def by simp
 qed
@@ -601,7 +601,7 @@
 using `finite I` proof induct
   case empty
   interpret finite_product_sigma_finite M "{}" by default simp
-  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> extreal"
+  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> ereal"
   show ?case
   proof (intro exI conjI ballI)
     have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
@@ -858,7 +858,7 @@
 qed
 
 lemma (in product_sigma_finite) product_positive_integral_setprod:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
   and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
@@ -874,14 +874,14 @@
     using insert by (auto intro!: setprod_cong)
   have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
     using sets_into_space insert
-    by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra
+    by (intro sigma_algebra.borel_measurable_ereal_setprod sigma_algebra_product_algebra
               measurable_comp[OF measurable_component_singleton, unfolded comp_def])
        auto
   then show ?case
     apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
-    apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc)
+    apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc)
     apply (subst I.positive_integral_cmult)
-    apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive)
+    apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive)
     done
 qed
 
@@ -890,8 +890,8 @@
   shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
 proof -
   interpret I: finite_product_sigma_finite M "{i}" by default simp
-  have *: "(\<lambda>x. extreal (f x)) \<in> borel_measurable (M i)"
-    "(\<lambda>x. extreal (- f x)) \<in> borel_measurable (M i)"
+  have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)"
+    "(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)"
     using assms by auto
   show ?thesis
     unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
@@ -965,17 +965,17 @@
   proof (unfold integrable_def, intro conjI)
     show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
       using borel by auto
-    have "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. extreal (abs (f i (x i)))) \<partial>P)"
-      by (simp add: setprod_extreal abs_setprod)
-    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. extreal (abs (f i x)) \<partial>M i))"
+    have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>P)"
+      by (simp add: setprod_ereal abs_setprod)
+    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))"
       using f by (subst product_positive_integral_setprod) auto
     also have "\<dots> < \<infinity>"
       using integrable[THEN M.integrable_abs]
       by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
-    finally show "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
-    have "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
+    finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
+    have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
       by (intro positive_integral_cong_pos) auto
-    then show "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
+    then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
   qed
   ultimately show ?thesis
     by (rule integrable_abs_iff[THEN iffD1])
--- a/src/HOL/Probability/Independent_Family.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Independent_Family.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -822,9 +822,9 @@
   assumes I: "I \<noteq> {}" "finite I"
   assumes rv: "\<And>i. random_variable (M' i) (X i)"
   shows "indep_vars M' X I \<longleftrightarrow>
-    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := extreal\<circ>distribution (X i) \<rparr>)).
+    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)).
       distribution (\<lambda>x. \<lambda>i\<in>I. X i x) A =
-      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := extreal\<circ>distribution (X i) \<rparr>)) A)"
+      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)) A)"
     (is "_ \<longleftrightarrow> (\<forall>X\<in>_. distribution ?D X = finite_measure.\<mu>' (Pi\<^isub>M I ?M) X)")
 proof -
   interpret M': prob_space "?M i" for i
@@ -832,7 +832,7 @@
   interpret P: finite_product_prob_space ?M I
     proof qed fact
 
-  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := extreal \<circ> distribution ?D \<rparr>"
+  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := ereal \<circ> distribution ?D \<rparr>"
   have "random_variable P.P ?D"
     using `finite I` rv by (intro random_variable_restrict) auto
   then interpret D: prob_space ?D'
@@ -938,24 +938,24 @@
 
 lemma (in prob_space) indep_var_distributionD:
   assumes indep: "indep_var S X T Y"
-  defines "P \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   assumes "A \<in> sets P"
   shows "joint_distribution X Y A = finite_measure.\<mu>' P A"
 proof -
   from indep have rvs: "random_variable S X" "random_variable T Y"
     by (blast dest: indep_var_rv1 indep_var_rv2)+
 
-  let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
-  let ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
+  let ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   interpret X: prob_space ?S by (rule distribution_prob_space) fact
   interpret Y: prob_space ?T by (rule distribution_prob_space) fact
   interpret XY: pair_prob_space ?S ?T by default
 
-  let ?J = "XY.P\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>"
+  let ?J = "XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>"
   interpret J: prob_space ?J
     by (rule joint_distribution_prob_space) (simp_all add: rvs)
 
-  have "finite_measure.\<mu>' (XY.P\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
+  have "finite_measure.\<mu>' (XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
   proof (rule prob_space_unique_Int_stable)
     show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P")
       by fact
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -563,7 +563,7 @@
             using `0 \<le> ?a` Q_sets J'.measure_space_1
             by (subst J'.positive_integral_add) auto
           finally show "?a / 2^(k+1) \<le> measure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
-            by (cases rule: extreal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
+            by (cases rule: ereal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
                (auto simp: field_simps)
         qed
         also have "\<dots> = measure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
@@ -712,7 +712,7 @@
       with `(\<Inter>i. A i) = {}` show False by auto
     qed
     ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
-      using LIMSEQ_extreal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
+      using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
   qed
 qed
 
@@ -812,7 +812,7 @@
   using assms
   unfolding \<mu>'_def M.\<mu>'_def
   by (subst measure_times[OF assms])
-     (auto simp: finite_measure_eq M.finite_measure_eq setprod_extreal)
+     (auto simp: finite_measure_eq M.finite_measure_eq setprod_ereal)
 
 lemma (in product_prob_space) finite_measure_infprod_emb_Pi:
   assumes J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> X j \<in> sets (M j)"
--- a/src/HOL/Probability/Information.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Information.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -203,7 +203,7 @@
 
   from real_RN_deriv[OF fms ac] guess D . note D = this
   with absolutely_continuous_AE[OF ms] ac
-  have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = extreal (D x)"
+  have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = ereal (D x)"
     by auto
 
   def f \<equiv> "\<lambda>x. if D x = 0 then 1 else 1 / D x"
@@ -215,7 +215,7 @@
     by (simp add: integral_cmult)
 
   { fix A assume "A \<in> sets M"
-    with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. extreal (D x) * indicator A x \<partial>M)"
+    with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. ereal (D x) * indicator A x \<partial>M)"
       by (auto intro!: positive_integral_cong_AE) }
   note D_density = this
 
@@ -231,12 +231,12 @@
   ultimately have M_int: "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
     by simp
 
-  have D_neg: "(\<integral>\<^isup>+ x. extreal (- D x) \<partial>M) = 0"
+  have D_neg: "(\<integral>\<^isup>+ x. ereal (- D x) \<partial>M) = 0"
     using D by (subst positive_integral_0_iff_AE) auto
 
-  have "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = \<nu> (space M)"
+  have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = \<nu> (space M)"
     using RN D by (auto intro!: positive_integral_cong_AE)
-  then have D_pos: "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = 1"
+  then have D_pos: "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = 1"
     using \<nu>.measure_space_1 by simp
 
   have "integrable M D"
@@ -271,16 +271,16 @@
 
       have "\<mu> {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
         using D(1) by auto
-      also have "\<dots> = (\<integral>\<^isup>+ x. extreal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
-        using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_extreal_def)
+      also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
+        using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def)
       also have "\<dots> = \<nu> {x\<in>space M. D x \<noteq> 0}"
         using D(1) D_density by auto
       also have "\<dots> = \<nu> (space M)"
         using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def)
       finally have "AE x. D x = 1"
         using D(1) \<nu>.measure_space_1 by (intro AE_I_eq_1) auto
-      then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. extreal (D x) * indicator A x\<partial>M)"
-        by (intro positive_integral_cong_AE) (auto simp: one_extreal_def[symmetric])
+      then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. ereal (D x) * indicator A x\<partial>M)"
+        by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
       also have "\<dots> = \<nu> A"
         using `A \<in> sets M` D_density by simp
       finally show False using `A \<in> sets M` `\<nu> A \<noteq> \<mu> A` by simp
@@ -293,7 +293,7 @@
       using D(2)
     proof (elim AE_mp, safe intro!: AE_I2)
       fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0"
-        and RN: "RN_deriv M \<nu> t = extreal (D t)"
+        and RN: "RN_deriv M \<nu> t = ereal (D t)"
         and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \<nu> t)"
 
       have "D t - 1 = D t - indicator ?D_set t"
@@ -311,7 +311,7 @@
     show "AE t. D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
       using D(2)
     proof (elim AE_mp, intro AE_I2 impI)
-      fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = extreal (D t)"
+      fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = ereal (D t)"
       show "D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
       proof cases
         assume asm: "D t \<noteq> 0"
@@ -347,7 +347,7 @@
       using eq by (intro measure_space_cong) auto
     show "absolutely_continuous \<nu>"
       unfolding absolutely_continuous_def using eq by auto
-    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: extreal)" by auto
+    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: ereal)" by auto
     fix A assume "A \<in> sets M"
     with eq show "\<nu> A = \<integral>\<^isup>+ x. 1 * indicator A x \<partial>M" by simp
   qed
@@ -468,17 +468,17 @@
 
 definition (in prob_space)
   "mutual_information b S T X Y =
-    KL_divergence b (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
-      (extreal\<circ>joint_distribution X Y)"
+    KL_divergence b (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
+      (ereal\<circ>joint_distribution X Y)"
 
 lemma (in information_space)
   fixes S T X Y
-  defines "P \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   shows "indep_var S X T Y \<longleftrightarrow>
     (random_variable S X \<and> random_variable T Y \<and>
-      measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y) \<and>
-      integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (extreal\<circ>joint_distribution X Y)) \<and>
+      measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y) \<and>
+      integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
+        (entropy_density b P (ereal\<circ>joint_distribution X Y)) \<and>
      mutual_information b S T X Y = 0)"
 proof safe
   assume indep: "indep_var S X T Y"
@@ -487,16 +487,16 @@
   then show "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
     by blast+
 
-  interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
     by (rule distribution_prob_space) fact
-  interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
     by (rule distribution_prob_space) fact
-  interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
   interpret XY: information_space XY.P b by default (rule b_gt_1)
 
-  let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
+  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
   { fix A assume "A \<in> sets XY.P"
-    then have "extreal (joint_distribution X Y A) = XY.\<mu> A"
+    then have "ereal (joint_distribution X Y A) = XY.\<mu> A"
       using indep_var_distributionD[OF indep]
       by (simp add: XY.P.finite_measure_eq) }
   note j_eq = this
@@ -504,31 +504,31 @@
   interpret J: prob_space ?J
     using j_eq by (intro XY.prob_space_cong) auto
 
-  have ac: "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
+  have ac: "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
     by (simp add: XY.absolutely_continuous_def j_eq)
-  then show "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
+  then show "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
     unfolding P_def .
 
-  have ed: "entropy_density b XY.P (extreal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
+  have ed: "entropy_density b XY.P (ereal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
     by (rule XY.measurable_entropy_density) (default | fact)+
 
-  have "AE x in XY.P. 1 = RN_deriv XY.P (extreal\<circ>joint_distribution X Y) x"
+  have "AE x in XY.P. 1 = RN_deriv XY.P (ereal\<circ>joint_distribution X Y) x"
   proof (rule XY.RN_deriv_unique[OF _ ac])
     show "measure_space ?J" by default
     fix A assume "A \<in> sets XY.P"
-    then show "(extreal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)"
+    then show "(ereal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)"
       by (simp add: j_eq)
   qed (insert XY.measurable_const[of 1 borel], auto)
-  then have ae_XY: "AE x in XY.P. entropy_density b XY.P (extreal\<circ>joint_distribution X Y) x = 0"
+  then have ae_XY: "AE x in XY.P. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
     by (elim XY.AE_mp) (simp add: entropy_density_def)
-  have ae_J: "AE x in ?J. entropy_density b XY.P (extreal\<circ>joint_distribution X Y) x = 0"
+  have ae_J: "AE x in ?J. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
   proof (rule XY.absolutely_continuous_AE)
     show "measure_space ?J" by default
     show "XY.absolutely_continuous (measure ?J)"
       using ac by simp
   qed (insert ae_XY, simp_all)
-  then show "integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (extreal\<circ>joint_distribution X Y))"
+  then show "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
+        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
     unfolding P_def
     using ed XY.measurable_const[of 0 borel]
     by (subst J.integrable_cong_AE) auto
@@ -540,30 +540,30 @@
   assume "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
   then have rvs: "random_variable S X" "random_variable T Y" by blast+
 
-  interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
     by (rule distribution_prob_space) fact
-  interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
     by (rule distribution_prob_space) fact
-  interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
   interpret XY: information_space XY.P b by default (rule b_gt_1)
 
-  let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
+  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
   interpret J: prob_space ?J
     using rvs by (intro joint_distribution_prob_space) auto
 
-  assume ac: "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
-  assume int: "integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (extreal\<circ>joint_distribution X Y))"
+  assume ac: "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
+  assume int: "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
+        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
   assume I_eq_0: "mutual_information b S T X Y = 0"
 
-  have eq: "\<forall>A\<in>sets XY.P. (extreal \<circ> joint_distribution X Y) A = XY.\<mu> A"
+  have eq: "\<forall>A\<in>sets XY.P. (ereal \<circ> joint_distribution X Y) A = XY.\<mu> A"
   proof (rule XY.KL_eq_0_imp)
     show "prob_space ?J" by default
-    show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
+    show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
       using ac by (simp add: P_def)
-    show "integrable ?J (entropy_density b XY.P (extreal\<circ>joint_distribution X Y))"
+    show "integrable ?J (entropy_density b XY.P (ereal\<circ>joint_distribution X Y))"
       using int by (simp add: P_def)
-    show "KL_divergence b XY.P (extreal\<circ>joint_distribution X Y) = 0"
+    show "KL_divergence b XY.P (ereal\<circ>joint_distribution X Y) = 0"
       using I_eq_0 unfolding mutual_information_def by (simp add: P_def)
   qed
 
@@ -587,13 +587,13 @@
         using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
     next
       fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
-      have "extreal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
-        extreal (joint_distribution X Y (A \<times> B))"
+      have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
+        ereal (joint_distribution X Y (A \<times> B))"
         unfolding distribution_def
-        by (intro arg_cong[where f="\<lambda>C. extreal (prob C)"]) auto
+        by (intro arg_cong[where f="\<lambda>C. ereal (prob C)"]) auto
       also have "\<dots> = XY.\<mu> (A \<times> B)"
         using ab eq by (auto simp: XY.finite_measure_eq)
-      also have "\<dots> = extreal (distribution X A) * extreal (distribution Y B)"
+      also have "\<dots> = ereal (distribution X A) * ereal (distribution Y B)"
         using ab by (simp add: XY.pair_measure_times)
       finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
         prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
@@ -605,10 +605,10 @@
 lemma (in information_space) mutual_information_commute_generic:
   assumes X: "random_variable S X" and Y: "random_variable T Y"
   assumes ac: "measure_space.absolutely_continuous
-    (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>) (extreal\<circ>joint_distribution X Y)"
+    (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>) (ereal\<circ>joint_distribution X Y)"
   shows "mutual_information b S T X Y = mutual_information b T S Y X"
 proof -
-  let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   interpret S: prob_space ?S using X by (rule distribution_prob_space)
   interpret T: prob_space ?T using Y by (rule distribution_prob_space)
   interpret P: pair_prob_space ?S ?T ..
@@ -617,13 +617,13 @@
     unfolding mutual_information_def
   proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1])
     show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
-      (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
+      (P.P \<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := ereal\<circ>joint_distribution Y X\<rparr>)"
       using X Y unfolding measurable_def
       unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable
       by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>'])
-    have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
+    have "prob_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
       using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
-    then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
+    then show "measure_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
       unfolding prob_space_def by simp
   qed auto
 qed
@@ -634,33 +634,33 @@
 abbreviation (in information_space)
   mutual_information_Pow ("\<I>'(_ ; _')") where
   "\<I>(X ; Y) \<equiv> mutual_information b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
 
 lemma (in prob_space) finite_variables_absolutely_continuous:
   assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
   shows "measure_space.absolutely_continuous
-    (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
-    (extreal\<circ>joint_distribution X Y)"
+    (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
+    (ereal\<circ>joint_distribution X Y)"
 proof -
-  interpret X: finite_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+  interpret X: finite_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
     using X by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  interpret Y: finite_prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
     using Y by (rule distribution_finite_prob_space)
   interpret XY: pair_finite_prob_space
-    "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr> measure := extreal\<circ>distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>"
+    "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr> measure := ereal\<circ>distribution Y\<rparr>" by default
+  interpret P: finite_prob_space "XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>"
     using assms by (auto intro!: joint_distribution_finite_prob_space)
   note rv = assms[THEN finite_random_variableD]
-  show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
+  show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
   proof (rule XY.absolutely_continuousI)
-    show "finite_measure_space (XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
+    show "finite_measure_space (XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
     fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
     then obtain a b where "x = (a, b)"
       and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
       by (cases x) (auto simp: space_pair_measure)
     with finite_distribution_order(5,6)[OF X Y]
-    show "(extreal \<circ> joint_distribution X Y) {x} = 0" by auto
+    show "(ereal \<circ> joint_distribution X Y) {x} = 0" by auto
   qed
 qed
 
@@ -676,16 +676,16 @@
   and mutual_information_positive_generic:
      "0 \<le> mutual_information b MX MY X Y" (is ?positive)
 proof -
-  interpret X: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+  interpret X: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
     using MX by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+  interpret Y: finite_prob_space "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
     using MY by (rule distribution_finite_prob_space)
-  interpret XY: pair_finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>" "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>"
+  interpret XY: pair_finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
+  interpret P: finite_prob_space "XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>"
     using assms by (auto intro!: joint_distribution_finite_prob_space)
 
-  have P_ms: "finite_measure_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
-  have P_ps: "finite_prob_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
+  have P_ms: "finite_measure_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
+  have P_ps: "finite_prob_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
 
   show ?sum
     unfolding Let_def mutual_information_def
@@ -739,14 +739,14 @@
 
 abbreviation (in information_space)
   entropy_Pow ("\<H>'(_')") where
-  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr> X"
+  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> X"
 
 lemma (in information_space) entropy_generic_eq:
   fixes X :: "'a \<Rightarrow> 'c"
   assumes MX: "finite_random_variable MX X"
   shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
 proof -
-  interpret MX: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+  interpret MX: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
     using MX by (rule distribution_finite_prob_space)
   let "?X x" = "distribution X {x}"
   let "?XX x y" = "joint_distribution X X {(x, y)}"
@@ -779,9 +779,9 @@
   assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
   shows "\<H>(X) = 0"
 proof -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal\<circ>distribution X\<rparr>"
+  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal\<circ>distribution X\<rparr>"
   note simple_function_imp_finite_random_variable[OF `simple_function M X`]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
+  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
   interpret X: finite_prob_space ?X by simp
   have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
     using X.measure_compl[of "{x}"] assms by auto
@@ -818,9 +818,9 @@
 
 lemma (in prob_space) measure'_translate:
   assumes X: "random_variable S X" and A: "A \<in> sets S"
-  shows "finite_measure.\<mu>' (S\<lparr> measure := extreal\<circ>distribution X \<rparr>) A = distribution X A"
+  shows "finite_measure.\<mu>' (S\<lparr> measure := ereal\<circ>distribution X \<rparr>) A = distribution X A"
 proof -
-  interpret S: prob_space "S\<lparr> measure := extreal\<circ>distribution X \<rparr>"
+  interpret S: prob_space "S\<lparr> measure := ereal\<circ>distribution X \<rparr>"
     using distribution_prob_space[OF X] .
   from A show "S.\<mu>' A = distribution X A"
     unfolding S.\<mu>'_def by (simp add: distribution_def_raw \<mu>'_def)
@@ -831,9 +831,9 @@
   assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
   shows "\<H>(X) = log b (real (card (X ` space M)))"
 proof -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := extreal\<circ>distribution X\<rparr>"
+  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := ereal\<circ>distribution X\<rparr>"
   note frv = simple_function_imp_finite_random_variable[OF X]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
+  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
   interpret X: finite_prob_space ?X by simp
   note rv = finite_random_variableD[OF frv]
   have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
@@ -917,9 +917,9 @@
 abbreviation (in information_space)
   conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
   "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr>
-    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = extreal\<circ>distribution Z \<rparr>
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr>
+    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = ereal\<circ>distribution Z \<rparr>
     X Y Z"
 
 lemma (in information_space) conditional_mutual_information_generic_eq:
@@ -1118,8 +1118,8 @@
 abbreviation (in information_space)
   conditional_entropy_Pow ("\<H>'(_ | _')") where
   "\<H>(X | Y) \<equiv> conditional_entropy b
-    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
+    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
+    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
 
 lemma (in information_space) conditional_entropy_positive:
   "simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -9,10 +9,10 @@
 imports Measure Borel_Space
 begin
 
-lemma real_extreal_1[simp]: "real (1::extreal) = 1"
-  unfolding one_extreal_def by simp
+lemma real_ereal_1[simp]: "real (1::ereal) = 1"
+  unfolding one_ereal_def by simp
 
-lemma extreal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::extreal)"
+lemma ereal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::ereal)"
   unfolding indicator_def by auto
 
 lemma tendsto_real_max:
@@ -150,7 +150,7 @@
 qed
 
 lemma (in sigma_algebra) simple_function_indicator_representation:
-  fixes f ::"'a \<Rightarrow> extreal"
+  fixes f ::"'a \<Rightarrow> ereal"
   assumes f: "simple_function M f" and x: "x \<in> space M"
   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
   (is "?l = ?r")
@@ -165,7 +165,7 @@
 qed
 
 lemma (in measure_space) simple_function_notspace:
-  "simple_function M (\<lambda>x. h x * indicator (- space M) x::extreal)" (is "simple_function M ?h")
+  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h")
 proof -
   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
   hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
@@ -212,7 +212,7 @@
   by (auto intro: borel_measurable_vimage)
 
 lemma (in sigma_algebra) simple_function_eq_borel_measurable:
-  fixes f :: "'a \<Rightarrow> extreal"
+  fixes f :: "'a \<Rightarrow> ereal"
   shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
   using simple_function_borel_measurable[of f]
     borel_measurable_simple_function[of f]
@@ -300,7 +300,7 @@
 
 lemma (in sigma_algebra)
   fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
-  shows simple_function_extreal[intro, simp]: "simple_function M (\<lambda>x. extreal (f x))"
+  shows simple_function_ereal[intro, simp]: "simple_function M (\<lambda>x. ereal (f x))"
   by (auto intro!: simple_function_compose1[OF sf])
 
 lemma (in sigma_algebra)
@@ -309,7 +309,7 @@
   by (auto intro!: simple_function_compose1[OF sf])
 
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
-  fixes u :: "'a \<Rightarrow> extreal"
+  fixes u :: "'a \<Rightarrow> ereal"
   assumes u: "u \<in> borel_measurable M"
   shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
              (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
@@ -331,7 +331,7 @@
     "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))"
     unfolding f_def by auto
 
-  let "?g j x" = "real (f x j) / 2^j :: extreal"
+  let "?g j x" = "real (f x j) / 2^j :: ereal"
   show ?thesis
   proof (intro exI[of _ ?g] conjI allI ballI)
     fix i
@@ -345,22 +345,22 @@
         by (rule finite_subset) auto
     qed
     then show "simple_function M (?g i)"
-      by (auto intro: simple_function_extreal simple_function_div)
+      by (auto intro: simple_function_ereal simple_function_div)
   next
     show "incseq ?g"
-    proof (intro incseq_extreal incseq_SucI le_funI)
+    proof (intro incseq_ereal incseq_SucI le_funI)
       fix x and i :: nat
       have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
       proof ((split split_if)+, intro conjI impI)
-        assume "extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+        assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
         then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)"
           by (cases "u x") (auto intro!: le_natfloor)
       next
-        assume "\<not> extreal (real i) \<le> u x" "extreal (real (Suc i)) \<le> u x"
+        assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x"
         then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i"
           by (cases "u x") auto
       next
-        assume "\<not> extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
+        assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
         have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2"
           by simp
         also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)"
@@ -380,7 +380,7 @@
     qed
   next
     fix x show "(SUP i. ?g i x) = max 0 (u x)"
-    proof (rule extreal_SUPI)
+    proof (rule ereal_SUPI)
       fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
         by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
                                      mult_nonpos_nonneg mult_nonneg_nonneg)
@@ -393,7 +393,7 @@
         case (real r)
         with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
         from real_arch_lt[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
-        then have "\<exists>p. max 0 (u x) = extreal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
+        then have "\<exists>p. max 0 (u x) = ereal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
         then guess p .. note ux = this
         obtain m :: nat where m: "p < real m" using real_arch_lt ..
         have "p \<le> r"
@@ -417,7 +417,7 @@
 qed
 
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
-  fixes u :: "'a \<Rightarrow> extreal"
+  fixes u :: "'a \<Rightarrow> ereal"
   assumes u: "u \<in> borel_measurable M"
   obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
     "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
@@ -454,7 +454,7 @@
 qed
 
 lemma (in measure_space) simple_function_restricted:
-  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
+  fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M"
   shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
     (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
 proof -
@@ -478,7 +478,7 @@
         using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
     next
       fix x
-      assume "indicator A x \<noteq> (0::extreal)"
+      assume "indicator A x \<noteq> (0::ereal)"
       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
@@ -527,7 +527,7 @@
   "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
 
 syntax
-  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
+  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
 
 translations
   "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
@@ -591,7 +591,7 @@
   hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
     unfolding simple_integral_def using f sets
     by (subst setsum_Sigma[symmetric])
-       (auto intro!: setsum_cong setsum_extreal_right_distrib)
+       (auto intro!: setsum_cong setsum_ereal_right_distrib)
   also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
   proof -
     have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
@@ -625,7 +625,7 @@
       simple_function_partition[OF f g]
       simple_function_partition[OF g f]
     by (subst (3) Int_commute)
-       (auto simp add: extreal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
+       (auto simp add: ereal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
 qed
 
 lemma (in measure_space) simple_integral_setsum[simp]:
@@ -650,8 +650,8 @@
   with assms show ?thesis
     unfolding simple_function_partition[OF mult f(1)]
               simple_function_partition[OF f(1) mult]
-    by (subst setsum_extreal_right_distrib)
-       (auto intro!: extreal_0_le_mult setsum_cong simp: mult_assoc)
+    by (subst setsum_ereal_right_distrib)
+       (auto intro!: ereal_0_le_mult setsum_cong simp: mult_assoc)
 qed
 
 lemma (in measure_space) simple_integral_mono_AE:
@@ -673,7 +673,7 @@
     proof (cases "f x \<le> g x")
       case True then show ?thesis
         using * assms(1,2)[THEN simple_functionD(2)]
-        by (auto intro!: extreal_mult_right_mono)
+        by (auto intro!: ereal_mult_right_mono)
     next
       case False
       obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
@@ -767,7 +767,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::extreal}" by auto
+  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::ereal}" by auto
   thus ?thesis
     using simple_integral_indicator[OF assms simple_function_const[of 1]]
     using sets_into_space[OF assms]
@@ -778,7 +778,7 @@
   assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
   shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
 proof -
-  have "AE x. indicator N x = (0 :: extreal)"
+  have "AE x. indicator N x = (0 :: ereal)"
     using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
   then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
     using assms apply (intro simple_integral_cong_AE) by auto
@@ -815,7 +815,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 extreal_mult_cancel_left by auto
+    unfolding ereal_mult_cancel_left by auto
 qed
 
 lemma (in measure_space) simple_integral_subalgebra:
@@ -872,7 +872,7 @@
   "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
 
 syntax
-  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
+  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
 
 translations
   "\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
@@ -917,8 +917,8 @@
     have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
     proof (intro SUP_PInfty)
       fix n :: nat
-      let ?y = "extreal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
-      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: extreal_divide_eq)
+      let ?y = "ereal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
+      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: ereal_divide_eq)
       then have "?g ?y \<in> ?A" by (rule g_in_A)
       have "real n \<le> ?y * \<mu> ?G"
         using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
@@ -1002,13 +1002,13 @@
   assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
   and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
   shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
-proof (rule extreal_le_mult_one_interval)
+proof (rule ereal_le_mult_one_interval)
   have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
     using f(3) by (auto intro!: le_SUPI2 positive_integral_positive)
   then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
   have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
     using u(3) by auto
-  fix a :: extreal assume "0 < a" "a < 1"
+  fix a :: ereal 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"
@@ -1043,7 +1043,7 @@
         assume "u x \<noteq> 0"
         with `a < 1` u_range[OF `x \<in> space M`]
         have "a * u x < 1 * u x"
-          by (intro extreal_mult_strict_right_mono) (auto simp: image_iff)
+          by (intro ereal_mult_strict_right_mono) (auto simp: image_iff)
         also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUPR_apply)
         finally obtain i where "a * u x < f i x" unfolding SUPR_def
           by (auto simp add: less_Sup_iff)
@@ -1056,18 +1056,18 @@
 
   have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
     unfolding simple_integral_indicator[OF B `simple_function M u`]
-  proof (subst SUPR_extreal_setsum, safe)
+  proof (subst SUPR_ereal_setsum, safe)
     fix x n assume "x \<in> space M"
     with u_range show "incseq (\<lambda>i. u x * \<mu> (?B' (u x) i))" "\<And>i. 0 \<le> u x * \<mu> (?B' (u x) i)"
-      using B_mono B_u by (auto intro!: measure_mono extreal_mult_left_mono incseq_SucI simp: extreal_zero_le_0_iff)
+      using B_mono B_u by (auto intro!: measure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
   next
     show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
       using measure_conv u_range B_u unfolding simple_integral_def
-      by (auto intro!: setsum_cong SUPR_extreal_cmult[symmetric])
+      by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric])
   qed
   moreover
   have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
-    apply (subst SUPR_extreal_cmult[symmetric])
+    apply (subst SUPR_ereal_cmult[symmetric])
   proof (safe intro!: SUP_mono bexI)
     fix i
     have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
@@ -1234,7 +1234,7 @@
   have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
     using u v `0 \<le> a`
     by (auto simp: incseq_Suc_iff le_fun_def
-             intro!: add_mono extreal_mult_left_mono simple_integral_mono)
+             intro!: add_mono ereal_mult_left_mono simple_integral_mono)
   have pos: "\<And>i. 0 \<le> integral\<^isup>S M (u i)" "\<And>i. 0 \<le> integral\<^isup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^isup>S M (u i)"
     using u v `0 \<le> a` by (auto simp: simple_integral_positive)
   { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
@@ -1245,26 +1245,26 @@
   proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
     show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
       using u v  `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
-      by (auto intro!: add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg)
+      by (auto intro!: add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg)
     { fix x
       { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
           by auto }
       then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
         using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
-        by (subst SUPR_extreal_cmult[symmetric, OF u(6) `0 \<le> a`])
-           (auto intro!: SUPR_extreal_add
-                 simp: incseq_Suc_iff le_fun_def add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) }
+        by (subst SUPR_ereal_cmult[symmetric, OF u(6) `0 \<le> a`])
+           (auto intro!: SUPR_ereal_add
+                 simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg) }
     then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
       unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
-      by (intro AE_I2) (auto split: split_max simp add: extreal_add_nonneg_nonneg)
+      by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg)
   qed
   also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
     using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
   finally have "(\<integral>\<^isup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^isup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+x. max 0 (g x) \<partial>M)"
     unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
     unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
-    apply (subst SUPR_extreal_cmult[symmetric, OF pos(1) `0 \<le> a`])
-    apply (subst SUPR_extreal_add[symmetric, OF inc not_MInf]) .
+    apply (subst SUPR_ereal_cmult[symmetric, OF pos(1) `0 \<le> a`])
+    apply (subst SUPR_ereal_add[symmetric, OF inc not_MInf]) .
   then show ?thesis by (simp add: positive_integral_max_0)
 qed
 
@@ -1273,7 +1273,7 @@
   shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
 proof -
   have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
-    by (auto split: split_max simp: extreal_zero_le_0_iff)
+    by (auto split: split_max simp: ereal_zero_le_0_iff)
   have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
     by (simp add: positive_integral_max_0)
   then show ?thesis
@@ -1302,7 +1302,7 @@
   shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
 proof -
   have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
-    using assms by (auto split: split_max simp: extreal_add_nonneg_nonneg)
+    using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg)
   have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
     by (simp add: positive_integral_max_0)
   also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
@@ -1325,7 +1325,7 @@
     case (insert i P)
     then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
       "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
-      by (auto intro!: borel_measurable_extreal_setsum setsum_nonneg)
+      by (auto intro!: borel_measurable_ereal_setsum setsum_nonneg)
     from positive_integral_add[OF this]
     show ?case using insert by auto
   qed simp
@@ -1342,10 +1342,10 @@
     using positive_integral_indicator by simp
   also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
     by (auto intro!: positive_integral_mono_AE
-      simp: indicator_def extreal_zero_le_0_iff)
+      simp: indicator_def ereal_zero_le_0_iff)
   also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
     using assms
-    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: extreal_zero_le_0_iff)
+    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: ereal_zero_le_0_iff)
   finally show ?thesis .
 qed
 
@@ -1375,10 +1375,10 @@
   shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
 proof -
   have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
-    using assms by (auto intro: extreal_diff_positive)
+    using assms by (auto intro: ereal_diff_positive)
   have pos_f: "AE x. 0 \<le> f x" using mono g by auto
-  { fix a b :: extreal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
-      by (cases rule: extreal2_cases[of a b]) auto }
+  { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
+      by (cases rule: ereal2_cases[of a b]) auto }
   note * = this
   then have "AE x. f x = f x - g x + g x"
     using mono positive_integral_noteq_infinite[OF g fin] assms by auto
@@ -1387,7 +1387,7 @@
     by (rule positive_integral_cong_AE)
   show ?thesis unfolding **
     using fin positive_integral_positive[of g]
-    by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
+    by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
 qed
 
 lemma (in measure_space) positive_integral_suminf:
@@ -1397,20 +1397,20 @@
   have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
     using assms by (auto simp: AE_all_countable)
   have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
-    using positive_integral_positive by (rule suminf_extreal_eq_SUPR)
+    using positive_integral_positive by (rule suminf_ereal_eq_SUPR)
   also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
     unfolding positive_integral_setsum[OF f] ..
   also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
     by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
        (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
   also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
-    by (intro positive_integral_cong_AE) (auto simp: suminf_extreal_eq_SUPR)
+    by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUPR)
   finally show ?thesis by simp
 qed
 
 text {* Fatou's lemma: convergence theorem on limes inferior *}
 lemma (in measure_space) positive_integral_lim_INF:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
   shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
 proof -
@@ -1435,7 +1435,7 @@
   show ?thesis
   proof
     have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
-      using u by (auto simp: extreal_zero_le_0_iff)
+      using u by (auto simp: ereal_zero_le_0_iff)
     then show "positive M' (measure M')" unfolding M'
       using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
     show "countably_additive M' (measure M')"
@@ -1449,7 +1449,7 @@
         by (simp add: positive_integral_suminf[OF _ pos, symmetric])
       also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
         by (intro positive_integral_cong_AE)
-           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_extreal)
+           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_ereal)
       also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
         unfolding suminf_indicator[OF disj] ..
       finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
@@ -1498,7 +1498,7 @@
     { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
       then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
       from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
-        by (subst setsum_extreal_right_distrib) (auto simp: ac_simps)
+        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
       also have "\<dots> = f x * G i x"
         by (simp add: indicator_def if_distrib setsum_cases)
       finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
@@ -1521,10 +1521,10 @@
   also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
     using f G' G(2)[THEN incseq_SucD] G
     by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
-       (auto simp: extreal_mult_left_mono le_fun_def extreal_zero_le_0_iff)
+       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
   also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
     by (intro positive_integral_cong_AE)
-       (auto simp add: SUPR_extreal_cmult split: split_max)
+       (auto simp add: SUPR_ereal_cmult split: split_max)
   finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
 qed
 
@@ -1541,16 +1541,16 @@
     with positive_integral_null_set[of ?A u] u
     show "integral\<^isup>P M u = 0" by (simp add: u_eq)
   next
-    { fix r :: extreal and n :: nat assume gt_1: "1 \<le> real n * r"
-      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_extreal_def)
-      then have "0 \<le> r" by (auto simp add: extreal_zero_less_0_iff) }
+    { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
+      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def)
+      then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) }
     note gt_1 = this
     assume *: "integral\<^isup>P M u = 0"
     let "?M n" = "{x \<in> space M. 1 \<le> real (n::nat) * u x}"
     have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
     proof -
       { fix n :: nat
-        from positive_integral_Markov_inequality[OF u pos, of ?A "extreal (real n)"]
+        from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"]
         have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
         moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
         ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
@@ -1566,7 +1566,7 @@
         fix n :: nat and x
         assume *: "1 \<le> real n * u x"
         also from gt_1[OF this] have "real n * u x \<le> real (Suc n) * u x"
-          using `0 \<le> u x` by (auto intro!: extreal_mult_right_mono)
+          using `0 \<le> u x` by (auto intro!: ereal_mult_right_mono)
         finally show "1 \<le> real (Suc n) * u x" by auto
       qed
     qed
@@ -1579,7 +1579,7 @@
         obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
         hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
         hence "1 \<le> real j * r" using real `0 < r` by auto
-        thus ?thesis using `0 < r` real by (auto simp: one_extreal_def)
+        thus ?thesis using `0 < r` real by (auto simp: one_ereal_def)
       qed (insert `0 < u x`, auto)
     qed auto
     finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
@@ -1618,7 +1618,7 @@
 proof -
   interpret R: measure_space ?R
     by (rule restricted_measure_space) fact
-  let "?I g x" = "g x * indicator A x :: extreal"
+  let "?I g x" = "g x * indicator A x :: ereal"
   show ?thesis
     unfolding positive_integral_def
     unfolding simple_function_restricted[OF A]
@@ -1675,15 +1675,15 @@
 
 definition integrable where
   "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
-    (\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
+    (\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
 
 lemma integrableD[dest]:
   assumes "integrable M f"
-  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
+  shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
   using assms unfolding integrable_def by auto
 
 definition lebesgue_integral_def:
-  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. extreal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. extreal (- f x) \<partial>M))"
+  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. ereal (- f x) \<partial>M))"
 
 syntax
   "_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
@@ -1694,13 +1694,13 @@
 lemma (in measure_space) integrableE:
   assumes "integrable M f"
   obtains r q where
-    "(\<integral>\<^isup>+x. extreal (f x)\<partial>M) = extreal r"
-    "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M) = extreal q"
+    "(\<integral>\<^isup>+x. ereal (f x)\<partial>M) = ereal r"
+    "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M) = ereal q"
     "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
   using assms unfolding integrable_def lebesgue_integral_def
-  using positive_integral_positive[of "\<lambda>x. extreal (f x)"]
-  using positive_integral_positive[of "\<lambda>x. extreal (-f x)"]
-  by (cases rule: extreal2_cases[of "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. extreal (f x)\<partial>M)"]) auto
+  using positive_integral_positive[of "\<lambda>x. ereal (f x)"]
+  using positive_integral_positive[of "\<lambda>x. ereal (-f x)"]
+  by (cases rule: ereal2_cases[of "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. ereal (f x)\<partial>M)"]) auto
 
 lemma (in measure_space) integral_cong:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
@@ -1722,8 +1722,8 @@
   assumes cong: "AE x. f x = g x"
   shows "integral\<^isup>L M f = integral\<^isup>L M g"
 proof -
-  have *: "AE x. extreal (f x) = extreal (g x)"
-    "AE x. extreal (- f x) = extreal (- g x)" using cong by auto
+  have *: "AE x. ereal (f x) = ereal (g x)"
+    "AE x. ereal (- f x) = ereal (- g x)" using cong by auto
   show ?thesis
     unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
 qed
@@ -1733,8 +1733,8 @@
   assumes "AE x. f x = g x"
   shows "integrable M f = integrable M g"
 proof -
-  have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (g x) \<partial>M)"
-    "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (- g x) \<partial>M)"
+  have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (g x) \<partial>M)"
+    "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (- g x) \<partial>M)"
     using assms by (auto intro!: positive_integral_cong_AE)
   with assms show ?thesis
     by (auto simp: integrable_def)
@@ -1746,11 +1746,11 @@
 
 lemma (in measure_space) integral_eq_positive_integral:
   assumes f: "\<And>x. 0 \<le> f x"
-  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
+  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
 proof -
-  { fix x have "max 0 (extreal (- f x)) = 0" using f[of x] by (simp split: split_max) }
-  then have "0 = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)" by simp
-  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
+  { fix x have "max 0 (ereal (- f x)) = 0" using f[of x] by (simp split: split_max) }
+  then have "0 = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)" by simp
+  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
   finally show ?thesis
     unfolding lebesgue_integral_def by simp
 qed
@@ -1762,7 +1762,7 @@
 proof -
   interpret T: measure_space M' by (rule measure_space_vimage[OF T])
   from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
+  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
     and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
     using f by (auto simp: comp_def)
   then show ?thesis
@@ -1777,7 +1777,7 @@
 proof -
   interpret T: measure_space M' by (rule measure_space_vimage[OF T])
   from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
+  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
     and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
     using f by (auto simp: comp_def)
   then show ?thesis
@@ -1795,27 +1795,27 @@
     and "integrable N g = integrable M (\<lambda>x. f x * g x)" (is ?integrable)
 proof -
   from f have ms: "measure_space (M\<lparr>measure := ?d\<rparr>)"
-    by (intro measure_space_density[where u="\<lambda>x. extreal (f x)"]) auto
+    by (intro measure_space_density[where u="\<lambda>x. ereal (f x)"]) auto
 
-  from ms density N have "(\<integral>\<^isup>+ x. g x \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (extreal (g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
+  from ms density N have "(\<integral>\<^isup>+ x. g x \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
     unfolding positive_integral_max_0
     by (intro measure_space.positive_integral_cong_measure) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (g x)) \<partial>M)"
+  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (g x)) \<partial>M)"
     using f g by (intro positive_integral_translated_density) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (f x * g x)) \<partial>M)"
+  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (f x * g x)) \<partial>M)"
     using f by (intro positive_integral_cong_AE)
-               (auto simp: extreal_max_0 zero_le_mult_iff split: split_max)
+               (auto simp: ereal_max_0 zero_le_mult_iff split: split_max)
   finally have pos: "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
     by (simp add: positive_integral_max_0)
   
-  from ms density N have "(\<integral>\<^isup>+ x. - (g x) \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (extreal (- g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
+  from ms density N have "(\<integral>\<^isup>+ x. - (g x) \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (- g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
     unfolding positive_integral_max_0
     by (intro measure_space.positive_integral_cong_measure) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (- g x)) \<partial>M)"
+  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (- g x)) \<partial>M)"
     using f g by (intro positive_integral_translated_density) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (- f x * g x)) \<partial>M)"
+  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (- f x * g x)) \<partial>M)"
     using f by (intro positive_integral_cong_AE)
-               (auto simp: extreal_max_0 mult_le_0_iff split: split_max)
+               (auto simp: ereal_max_0 mult_le_0_iff split: split_max)
   finally have neg: "(\<integral>\<^isup>+ x. - g x \<partial>N) = (\<integral>\<^isup>+ x. - (f x * g x) \<partial>M)"
     by (simp add: positive_integral_max_0)
 
@@ -1846,10 +1846,10 @@
   and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
   shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
 proof -
-  let "?f x" = "max 0 (extreal (f x))"
-  let "?mf x" = "max 0 (extreal (- f x))"
-  let "?u x" = "max 0 (extreal (u x))"
-  let "?v x" = "max 0 (extreal (v x))"
+  let "?f x" = "max 0 (ereal (f x))"
+  let "?mf x" = "max 0 (ereal (- f x))"
+  let "?u x" = "max 0 (ereal (u x))"
+  let "?v x" = "max 0 (ereal (v x))"
 
   from borel_measurable_diff[of u v] integrable
   have f_borel: "?f \<in> borel_measurable M" and
@@ -1859,9 +1859,9 @@
     "f \<in> borel_measurable M"
     by (auto simp: f_def[symmetric] integrable_def)
 
-  have "(\<integral>\<^isup>+ x. extreal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
+  have "(\<integral>\<^isup>+ x. ereal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
     using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
-  moreover have "(\<integral>\<^isup>+ x. extreal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
+  moreover have "(\<integral>\<^isup>+ x. ereal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
     using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
   ultimately show f: "integrable M f"
     using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
@@ -1886,22 +1886,22 @@
   shows "integrable M (\<lambda>t. a * f t + g t)"
   and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
 proof -
-  let "?f x" = "max 0 (extreal (f x))"
-  let "?g x" = "max 0 (extreal (g x))"
-  let "?mf x" = "max 0 (extreal (- f x))"
-  let "?mg x" = "max 0 (extreal (- g x))"
+  let "?f x" = "max 0 (ereal (f x))"
+  let "?g x" = "max 0 (ereal (g x))"
+  let "?mf x" = "max 0 (ereal (- f x))"
+  let "?mg x" = "max 0 (ereal (- g x))"
   let "?p t" = "max 0 (a * f t) + max 0 (g t)"
   let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
 
   from assms have linear:
-    "(\<integral>\<^isup>+ x. extreal a * ?f x + ?g x \<partial>M) = extreal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
-    "(\<integral>\<^isup>+ x. extreal a * ?mf x + ?mg x \<partial>M) = extreal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
+    "(\<integral>\<^isup>+ x. ereal a * ?f x + ?g x \<partial>M) = ereal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
+    "(\<integral>\<^isup>+ x. ereal a * ?mf x + ?mg x \<partial>M) = ereal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
     by (auto intro!: positive_integral_linear simp: integrable_def)
 
-  have *: "(\<integral>\<^isup>+x. extreal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- ?n x) \<partial>M) = 0"
+  have *: "(\<integral>\<^isup>+x. ereal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- ?n x) \<partial>M) = 0"
     using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
-  have **: "\<And>x. extreal a * ?f x + ?g x = max 0 (extreal (?p x))"
-           "\<And>x. extreal a * ?mf x + ?mg x = max 0 (extreal (?n x))"
+  have **: "\<And>x. ereal a * ?f x + ?g x = max 0 (ereal (?p x))"
+           "\<And>x. ereal a * ?mf x + ?mg x = max 0 (ereal (?n x))"
     using `0 \<le> a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff)
 
   have "integrable M ?p" "integrable M ?n"
@@ -1958,12 +1958,12 @@
   and mono: "AE t. f t \<le> g t"
   shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
 proof -
-  have "AE x. extreal (f x) \<le> extreal (g x)"
+  have "AE x. ereal (f x) \<le> ereal (g x)"
     using mono by auto
-  moreover have "AE x. extreal (- g x) \<le> extreal (- f x)"
+  moreover have "AE x. ereal (- g x) \<le> ereal (- f x)"
     using mono by auto
   ultimately show ?thesis using fg
-    by (auto intro!: add_mono positive_integral_mono_AE real_of_extreal_positive_mono
+    by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
              simp: positive_integral_positive lebesgue_integral_def diff_minus)
 qed
 
@@ -1986,9 +1986,9 @@
   and "integrable M (indicator A)" (is ?able)
 proof -
   from `A \<in> sets M` have *:
-    "\<And>x. extreal (indicator A x) = indicator A x"
-    "(\<integral>\<^isup>+x. extreal (- indicator A x) \<partial>M) = 0"
-    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_extreal_def)
+    "\<And>x. ereal (indicator A x) = indicator A x"
+    "(\<integral>\<^isup>+x. ereal (- indicator A x) \<partial>M) = 0"
+    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_ereal_def)
   show ?int ?able
     using assms unfolding lebesgue_integral_def integrable_def
     by (auto simp: * positive_integral_indicator borel_measurable_indicator)
@@ -2027,8 +2027,8 @@
   assumes "integrable M f"
   shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
 proof -
-  from assms have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>)\<partial>M) = 0"
-    "\<And>x. extreal \<bar>f x\<bar> = max 0 (extreal (f x)) + max 0 (extreal (- f x))"
+  from assms have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
+    "\<And>x. ereal \<bar>f x\<bar> = max 0 (ereal (f x)) + max 0 (ereal (- f x))"
     by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
   with assms show ?thesis
     by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
@@ -2041,8 +2041,8 @@
     and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
 proof -
   interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
-  have "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M)"
-       "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)"
+  have "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M)"
+       "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)"
     using borel by (auto intro!: positive_integral_subalgebra N sa)
   moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
     using assms unfolding measurable_def by auto
@@ -2057,21 +2057,21 @@
   assumes borel: "g \<in> borel_measurable M"
   shows "integrable M g"
 proof -
-  have "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal \<bar>g x\<bar> \<partial>M)"
+  have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
     by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
     using f by (auto intro!: positive_integral_mono)
   also have "\<dots> < \<infinity>"
     using `integrable M f` unfolding integrable_def by auto
-  finally have pos: "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) < \<infinity>" .
+  finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
 
-  have "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal (\<bar>g x\<bar>) \<partial>M)"
+  have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
     by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
     using f by (auto intro!: positive_integral_mono)
   also have "\<dots> < \<infinity>"
     using `integrable M f` unfolding integrable_def by auto
-  finally have neg: "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) < \<infinity>" .
+  finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
 
   from neg pos borel show ?thesis
     unfolding integrable_def by auto
@@ -2143,31 +2143,31 @@
       by (simp add: mono_def incseq_def) }
   note pos_u = this
 
-  have SUP_F: "\<And>x. (SUP n. extreal (f n x)) = extreal (u x)"
+  have SUP_F: "\<And>x. (SUP n. ereal (f n x)) = ereal (u x)"
     unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
 
-  have borel_f: "\<And>i. (\<lambda>x. extreal (f i x)) \<in> borel_measurable M"
+  have borel_f: "\<And>i. (\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
     using i unfolding integrable_def by auto
-  hence "(\<lambda>x. SUP i. extreal (f i x)) \<in> borel_measurable M"
+  hence "(\<lambda>x. SUP i. ereal (f i x)) \<in> borel_measurable M"
     by auto
   hence borel_u: "u \<in> borel_measurable M"
-    by (auto simp: borel_measurable_extreal_iff SUP_F)
+    by (auto simp: borel_measurable_ereal_iff SUP_F)
 
-  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. extreal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- u x) \<partial>M) = 0"
+  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. ereal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- u x) \<partial>M) = 0"
     using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
 
-  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M) = extreal (integral\<^isup>L M (f n))"
-    using i positive_integral_positive by (auto simp: extreal_real lebesgue_integral_def integrable_def)
+  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M) = ereal (integral\<^isup>L M (f n))"
+    using i positive_integral_positive by (auto simp: ereal_real lebesgue_integral_def integrable_def)
 
   have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
     using pos i by (auto simp: integral_positive)
   hence "0 \<le> x"
     using LIMSEQ_le_const[OF ilim, of 0] by auto
 
-  from mono pos i have pI: "(\<integral>\<^isup>+ x. extreal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M))"
+  from mono pos i have pI: "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M))"
     by (auto intro!: positive_integral_monotone_convergence_SUP
       simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
-  also have "\<dots> = extreal x" unfolding integral_eq
+  also have "\<dots> = ereal x" unfolding integral_eq
   proof (rule SUP_eq_LIMSEQ[THEN iffD2])
     show "mono (\<lambda>n. integral\<^isup>L M (f n))"
       using mono i by (auto simp: mono_def intro!: integral_mono)
@@ -2205,15 +2205,15 @@
   assumes "integrable M f"
   shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
 proof -
-  have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>) \<partial>M) = 0"
+  have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
     using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
   have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
-  hence "(\<lambda>x. extreal (\<bar>f x\<bar>)) \<in> borel_measurable M"
-    "(\<integral>\<^isup>+ x. extreal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
+  hence "(\<lambda>x. ereal (\<bar>f x\<bar>)) \<in> borel_measurable M"
+    "(\<integral>\<^isup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
   from positive_integral_0_iff[OF this(1)] this(2)
   show ?thesis unfolding lebesgue_integral_def *
-    using positive_integral_positive[of "\<lambda>x. extreal \<bar>f x\<bar>"]
-    by (auto simp add: real_of_extreal_eq_0)
+    using positive_integral_positive[of "\<lambda>x. ereal \<bar>f x\<bar>"]
+    by (auto simp add: real_of_ereal_eq_0)
 qed
 
 lemma (in measure_space) positive_integral_PInf:
@@ -2255,17 +2255,17 @@
 lemma (in measure_space) integral_real:
   "AE x. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
   using assms unfolding lebesgue_integral_def
-  by (subst (1 2) positive_integral_cong_AE) (auto simp add: extreal_real)
+  by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real)
 
 lemma (in finite_measure) lebesgue_integral_const[simp]:
   shows "integrable M (\<lambda>x. a)"
   and  "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
 proof -
   { fix a :: real assume "0 \<le> a"
-    then have "(\<integral>\<^isup>+ x. extreal a \<partial>M) = extreal a * \<mu> (space M)"
+    then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * \<mu> (space M)"
       by (subst positive_integral_const) auto
     moreover
-    from `0 \<le> a` have "(\<integral>\<^isup>+ x. extreal (-a) \<partial>M) = 0"
+    from `0 \<le> a` have "(\<integral>\<^isup>+ x. ereal (-a) \<partial>M) = 0"
       by (subst positive_integral_0_iff_AE) auto
     ultimately have "integrable M (\<lambda>x. a)" by (auto simp: integrable_def) }
   note * = this
@@ -2282,7 +2282,7 @@
 qed
 
 lemma indicator_less[simp]:
-  "indicator A x \<le> (indicator B x::extreal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
+  "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
   by (simp add: indicator_def not_le)
 
 lemma (in finite_measure) integral_less_AE:
@@ -2365,21 +2365,21 @@
     finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
   note diff_less_2w = this
 
-  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. extreal (?diff n x) \<partial>M) =
-    (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. ereal (?diff n x) \<partial>M) =
+    (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
     using diff w diff_less_2w w_pos
     by (subst positive_integral_diff[symmetric])
        (auto simp: integrable_def intro!: positive_integral_cong)
 
   have "integrable M (\<lambda>x. 2 * w x)"
     using w by (auto intro: integral_cmult)
-  hence I2w_fin: "(\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
-    borel_2w: "(\<lambda>x. extreal (2 * w x)) \<in> borel_measurable M"
+  hence I2w_fin: "(\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
+    borel_2w: "(\<lambda>x. ereal (2 * w x)) \<in> borel_measurable M"
     unfolding integrable_def by auto
 
-  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
+  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
   proof cases
-    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
+    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
     { fix n
       have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
         using diff_less_2w[of _ n] unfolding positive_integral_max_0
@@ -2388,53 +2388,53 @@
         using positive_integral_positive[of ?f'] eq_0 by auto }
     then show ?thesis by (simp add: Limsup_const)
   next
-    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
-    have "0 = limsup (\<lambda>n. 0 :: extreal)" by (simp add: Limsup_const)
-    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
+    have "0 = limsup (\<lambda>n. 0 :: ereal)" by (simp add: Limsup_const)
+    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
       by (intro limsup_mono positive_integral_positive)
-    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)" .
-    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (extreal (?diff n x))) \<partial>M)"
+    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
+    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
     proof (rule positive_integral_cong)
       fix x assume x: "x \<in> space M"
-      show "max 0 (extreal (2 * w x)) = liminf (\<lambda>n. max 0 (extreal (?diff n x)))"
-        unfolding extreal_max_0
-      proof (rule lim_imp_Liminf[symmetric], unfold lim_extreal)
+      show "max 0 (ereal (2 * w x)) = liminf (\<lambda>n. max 0 (ereal (?diff n x)))"
+        unfolding ereal_max_0
+      proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal)
         have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
           using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
         then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
-          by (auto intro!: tendsto_real_max simp add: lim_extreal)
+          by (auto intro!: tendsto_real_max simp add: lim_ereal)
       qed (rule trivial_limit_sequentially)
     qed
-    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (extreal (?diff n x)) \<partial>M)"
+    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
       using u'_borel w u unfolding integrable_def
       by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
-    also have "\<dots> = (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) -
-        limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
+    also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
+        limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
       unfolding PI_diff positive_integral_max_0
-      using positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"]
-      by (subst liminf_extreal_cminus) auto
+      using positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"]
+      by (subst liminf_ereal_cminus) auto
     finally show ?thesis
-      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"] pos
+      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"] pos
       unfolding positive_integral_max_0
-      by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"])
+      by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"])
          auto
   qed
 
   have "liminf ?f \<le> limsup ?f"
-    by (intro extreal_Liminf_le_Limsup trivial_limit_sequentially)
+    by (intro ereal_Liminf_le_Limsup trivial_limit_sequentially)
   moreover
-  { have "0 = liminf (\<lambda>n. 0 :: extreal)" by (simp add: Liminf_const)
+  { have "0 = liminf (\<lambda>n. 0 :: ereal)" by (simp add: Liminf_const)
     also have "\<dots> \<le> liminf ?f"
       by (intro liminf_mono positive_integral_positive)
     finally have "0 \<le> liminf ?f" . }
-  ultimately have liminf_limsup_eq: "liminf ?f = extreal 0" "limsup ?f = extreal 0"
+  ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
     using `limsup ?f = 0` by auto
-  have "\<And>n. (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = extreal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
+  have "\<And>n. (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = ereal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
     using diff positive_integral_positive
-    by (subst integral_eq_positive_integral) (auto simp: extreal_real integrable_def)
+    by (subst integral_eq_positive_integral) (auto simp: ereal_real integrable_def)
   then show ?lim_diff
-    using extreal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
-    by (simp add: lim_extreal)
+    using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
+    by (simp add: lim_ereal)
 
   show ?lim
   proof (rule LIMSEQ_I)
@@ -2554,7 +2554,7 @@
     also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
       using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
     finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
-      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_extreal_pos measurable_sets) }
+      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) }
   note int_abs_F = this
 
   have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
@@ -2618,15 +2618,15 @@
   and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
 proof -
   have *:
-    "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (f x)) * \<mu> {x})"
-    "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (- f x)) * \<mu> {x})"
+    "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (f x)) * \<mu> {x})"
+    "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (- f x)) * \<mu> {x})"
     by (simp_all add: positive_integral_finite_eq_setsum)
   then show "integrable M f" using finite_space finite_measure
     by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
              split: split_max)
   show ?I using finite_measure *
     apply (simp add: positive_integral_max_0 lebesgue_integral_def)
-    apply (subst (1 2) setsum_real_of_extreal[symmetric])
+    apply (subst (1 2) setsum_real_of_ereal[symmetric])
     apply (simp_all split: split_max add: setsum_subtractf[symmetric])
     apply (intro setsum_cong[OF refl])
     apply (simp split: split_max)
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -55,7 +55,7 @@
 definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
   "lebesgue = \<lparr> space = UNIV,
     sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n},
-    measure = \<lambda>A. SUP n. extreal (integral (cube n) (indicator A)) \<rparr>"
+    measure = \<lambda>A. SUP n. ereal (integral (cube n) (indicator A)) \<rparr>"
 
 lemma space_lebesgue[simp]: "space lebesgue = UNIV"
   unfolding lebesgue_def by simp
@@ -141,15 +141,15 @@
       by (auto dest: lebesgueD)
     show "(\<Sum>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
     proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI)
-      fix i n show "extreal (?m n i) \<le> extreal (?m (Suc n) i)"
+      fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)"
         using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
     next
-      fix i n show "0 \<le> extreal (?m n i)"
+      fix i n show "0 \<le> ereal (?m n i)"
         using rA unfolding lebesgue_def
         by (auto intro!: le_SUPI2 integral_nonneg)
     next
-      show "(SUP n. \<Sum>i. extreal (?m n i)) = (SUP n. extreal (?M n UNIV))"
-      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_extreal[THEN iffD2] sums_def[THEN iffD2])
+      show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))"
+      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2])
         fix n
         have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
         from lebesgueD[OF this]
@@ -235,7 +235,7 @@
 
 lemma lmeasure_iff_LIMSEQ:
   assumes "A \<in> sets lebesgue" "0 \<le> m"
-  shows "lebesgue.\<mu> A = extreal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
+  shows "lebesgue.\<mu> A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
 proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ)
   show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
     using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
@@ -261,7 +261,7 @@
 
 lemma lmeasure_finite_has_integral:
   fixes s :: "'a::ordered_euclidean_space set"
-  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = extreal m" "0 \<le> m"
+  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = ereal m" "0 \<le> m"
   shows "(indicator s has_integral m) UNIV"
 proof -
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
@@ -324,7 +324,7 @@
 qed
 
 lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
-  shows "lebesgue.\<mu> s = extreal m"
+  shows "lebesgue.\<mu> s = ereal m"
 proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
   show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
@@ -349,28 +349,28 @@
 qed
 
 lemma has_integral_iff_lmeasure:
-  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m)"
+  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m)"
 proof
   assume "(indicator A has_integral m) UNIV"
   with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
-  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
+  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
     by (auto intro: has_integral_nonneg)
 next
-  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
+  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
   then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
 qed
 
 lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
-  shows "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))"
+  shows "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))"
   using assms unfolding integrable_on_def
 proof safe
   fix y :: real assume "(indicator s has_integral y) UNIV"
   from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
-  show "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))" by simp
+  show "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))" by simp
 qed
 
 lemma lebesgue_simple_function_indicator:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f:"simple_function lebesgue f"
   shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
   by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
@@ -427,14 +427,14 @@
         using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
       finally show ?thesis .
     qed }
-  ultimately show "extreal (real n) \<le> extreal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
+  ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
     using integral_const DIM_positive[where 'a='a]
     by (auto simp: cube_def content_closed_interval_cases setprod_constant)
 qed simp
 
 lemma
   fixes a b ::"'a::ordered_euclidean_space"
-  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = extreal (content {a..b})"
+  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = ereal (content {a..b})"
 proof -
   have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
     unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw)
@@ -462,7 +462,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanAtMost[simp]:
-    "lebesgue.\<mu> {a <.. b} = extreal (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a <.. b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
@@ -474,7 +474,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_atLeastLessThan[simp]:
-    "lebesgue.\<mu> {a ..< b} = extreal (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a ..< b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
@@ -486,7 +486,7 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanLessThan[simp]:
-    "lebesgue.\<mu> {a <..< b} = extreal (if a \<le> b then b - a else 0)"
+    "lebesgue.\<mu> {a <..< b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
   then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
@@ -553,7 +553,7 @@
 qed simp
 
 lemma simple_function_has_integral:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f:"simple_function lebesgue f"
   and f':"range f \<subseteq> {0..<\<infinity>}"
   and om:"\<And>x. x \<in> range f \<Longrightarrow> lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
@@ -563,16 +563,16 @@
   let "?M x" = "lebesgue.\<mu> (f -` {x} \<inter> UNIV)"
   let "?F x" = "indicator (f -` {x})"
   { fix x y assume "y \<in> range f"
-    from subsetD[OF f' this] have "y * ?F y x = extreal (real y * ?F y x)"
-      by (cases rule: extreal2_cases[of y "?F y x"])
-         (auto simp: indicator_def one_extreal_def split: split_if_asm) }
+    from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)"
+      by (cases rule: ereal2_cases[of y "?F y x"])
+         (auto simp: indicator_def one_ereal_def split: split_if_asm) }
   moreover
   { fix x assume x: "x\<in>range f"
     have "x * ?M x = real x * real (?M x)"
     proof cases
       assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
       with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis
-        by (cases rule: extreal2_cases[of x "?M x"]) auto
+        by (cases rule: ereal2_cases[of x "?M x"]) auto
     qed simp }
   ultimately
   have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow>
@@ -580,12 +580,12 @@
     by simp
   also have \<dots>
   proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
-               real_of_extreal_pos lebesgue.positive_measure ballI)
+               real_of_ereal_pos lebesgue.positive_measure ballI)
     show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue" "\<And>y. f -` {y} \<inter> UNIV \<in> sets lebesgue"
       using lebesgue.simple_functionD[OF f] by auto
     fix y assume "real y \<noteq> 0" "y \<in> range f"
-    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = extreal (real (?M y))"
-      by (auto simp: extreal_real)
+    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = ereal (real (?M y))"
+      by (auto simp: ereal_real)
   qed
   finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .
 qed fact
@@ -595,7 +595,7 @@
   using assms by auto
 
 lemma simple_function_has_integral':
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"
   and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
@@ -629,7 +629,7 @@
 qed
 
 lemma positive_integral_has_integral:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
 proof -
@@ -648,23 +648,23 @@
   note u_fin = this
   then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
     by (rule simple_function_has_integral'[OF u(1,5)])
-  have "\<forall>x. \<exists>r\<ge>0. f x = extreal r"
+  have "\<forall>x. \<exists>r\<ge>0. f x = ereal r"
   proof
     fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)
-    then show "\<exists>r\<ge>0. f x = extreal r" by (cases "f x") auto
+    then show "\<exists>r\<ge>0. f x = ereal r" by (cases "f x") auto
   qed
-  from choice[OF this] obtain f' where f': "f = (\<lambda>x. extreal (f' x))" "\<And>x. 0 \<le> f' x" by auto
+  from choice[OF this] obtain f' where f': "f = (\<lambda>x. ereal (f' x))" "\<And>x. 0 \<le> f' x" by auto
 
-  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
   proof
-    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
+    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
     proof (intro choice allI)
       fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis
-      then show "\<exists>r\<ge>0. u i x = extreal r" using u(5)[of i x] by (cases "u i x") auto
+      then show "\<exists>r\<ge>0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto
     qed
   qed
   from choice[OF this] obtain u' where
-      u': "u = (\<lambda>i x. extreal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
+      u': "u = (\<lambda>i x. ereal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
 
   have convergent: "f' integrable_on UNIV \<and>
     (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"
@@ -672,7 +672,7 @@
     show int: "\<And>k. (u' k) integrable_on UNIV"
       using u_int unfolding integrable_on_def u' by auto
     show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)
-      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_extreal_positive_mono)
+      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono)
     show "\<And>x. (\<lambda>k. u' k x) ----> f' x"
       using SUP_eq u(2)
       by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)
@@ -686,16 +686,16 @@
       also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
         using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp
       also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
-        by (auto intro!: real_of_extreal_positive_mono lebesgue.positive_integral_positive
+        by (auto intro!: real_of_ereal_positive_mono lebesgue.positive_integral_positive
              lebesgue.positive_integral_mono le_SUPI simp: SUP_eq[symmetric])
       finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
     qed
   qed
 
-  have "integral\<^isup>P lebesgue f = extreal (integral UNIV f')"
+  have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
   proof (rule tendsto_unique[OF trivial_limit_sequentially])
     have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
-      unfolding u_eq by (intro LIMSEQ_extreal_SUPR lebesgue.incseq_positive_integral u)
+      unfolding u_eq by (intro LIMSEQ_ereal_SUPR lebesgue.incseq_positive_integral u)
     also note lebesgue.positive_integral_monotone_convergence_SUP
       [OF u(2)  lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric]
     finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
@@ -704,12 +704,12 @@
     { fix k
       have "0 \<le> integral\<^isup>S lebesgue (u k)"
         using u by (auto intro!: lebesgue.simple_integral_positive)
-      then have "integral\<^isup>S lebesgue (u k) = extreal (real (integral\<^isup>S lebesgue (u k)))"
-        using u_fin by (auto simp: extreal_real) }
+      then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"
+        using u_fin by (auto simp: ereal_real) }
     note * = this
-    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> extreal (integral UNIV f')"
+    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
       using convergent using u_int[THEN integral_unique, symmetric]
-      by (subst *) (simp add: lim_extreal u')
+      by (subst *) (simp add: lim_ereal u')
   qed
   then show ?thesis using convergent by (simp add: f' integrable_integral)
 qed
@@ -719,9 +719,9 @@
   assumes f: "integrable lebesgue f"
   shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
 proof -
-  let ?n = "\<lambda>x. real (extreal (max 0 (- f x)))" and ?p = "\<lambda>x. real (extreal (max 0 (f x)))"
-  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: extreal_max)
-  { fix f have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. extreal (max 0 (f x)) \<partial>lebesgue)"
+  let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"
+  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)
+  { fix f have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
       by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) }
   note eq = this
   show ?thesis
@@ -731,8 +731,8 @@
     unfolding eq[of f] eq[of "\<lambda>x. - f x"]
     apply (safe intro!: positive_integral_has_integral)
     using integrableD[OF f]
-    by (auto simp: zero_extreal_def[symmetric] positive_integral_max_0  split: split_max
-             intro!: lebesgue.measurable_If lebesgue.borel_measurable_extreal)
+    by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0  split: split_max
+             intro!: lebesgue.measurable_If lebesgue.borel_measurable_ereal)
 qed
 
 lemma lebesgue_positive_integral_eq_borel:
@@ -892,7 +892,7 @@
         by (simp add: interval_ne_empty eucl_le[where 'a='a])
       then have "lborel.\<mu> X = (\<Prod>x<DIM('a). lborel.\<mu> {a $$ x .. b $$ x})"
         by (auto simp: content_closed_interval eucl_le[where 'a='a]
-                 intro!: setprod_extreal[symmetric])
+                 intro!: setprod_ereal[symmetric])
       also have "\<dots> = measure ?P (?T X)"
         unfolding * by (subst lborel_space.measure_times) auto
       finally show ?thesis .
@@ -917,7 +917,7 @@
   using lborel_eq_lborel_space[OF A] by simp
 
 lemma borel_fubini_positiv_integral:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable borel"
   shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
 proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _])
@@ -982,7 +982,7 @@
 lemma lebesgue_real_affine:
   fixes X :: "real set"
   assumes "X \<in> sets borel" and "c \<noteq> 0"
-  shows "measure lebesgue X = extreal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
+  shows "measure lebesgue X = ereal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
     (is "_ = ?\<nu> X")
 proof -
   let ?E = "\<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::real .. b})\<rparr> :: real algebra"
@@ -1010,7 +1010,7 @@
     show "measure_space (lborel\<lparr>measure := ?\<nu>\<rparr>)"
     proof
       show "positive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
-        by (auto simp: positive_def intro!: extreal_0_le_mult borel.borel_measurable_sets)
+        by (auto simp: positive_def intro!: ereal_0_le_mult borel.borel_measurable_sets)
       show "countably_additive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
       proof (simp add: countably_additive_def, safe)
         fix A :: "nat \<Rightarrow> real set" assume A: "range A \<subseteq> sets borel" "disjoint_family A"
@@ -1026,7 +1026,7 @@
             by (rule disjoint_family_on_bisimulation) auto
         qed
         with Ai show "(\<Sum>n. ?\<nu> (A n)) = ?\<nu> (UNION UNIV A)"
-          by (subst suminf_cmult_extreal)
+          by (subst suminf_cmult_ereal)
              (auto simp: vimage_UN borel.borel_measurable_sets)
       qed
     qed
--- a/src/HOL/Probability/Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -150,7 +150,7 @@
   finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
   then show ?thesis
     using fin `0 \<le> \<mu> s`
-    unfolding extreal_eq_minus_iff by (auto simp: ac_simps)
+    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
 qed
 
 lemma (in measure_space) measure_Diff:
@@ -164,7 +164,7 @@
   also have "\<dots> = \<mu> (A - B) + \<mu> B"
     using measurable by (subst measure_additive[symmetric]) auto
   finally show "\<mu> (A - B) = \<mu> A - \<mu> B"
-    unfolding extreal_eq_minus_iff
+    unfolding ereal_eq_minus_iff
     using finite `0 \<le> \<mu> B` by auto
 qed
 
@@ -207,7 +207,7 @@
   have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
     by (blast intro: range_subsetD [OF A])
   have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))"
-    using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric])
+    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
   also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)"
     by (rule measure_countably_additive)
        (auto simp add: disjoint_family_Suc ASuc A1 A2)
@@ -244,7 +244,7 @@
 lemma (in measure_space) continuity_from_below_Lim:
   assumes A: "range A \<subseteq> sets M" "incseq A"
   shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)"
-  using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A]
+  using LIMSEQ_ereal_SUPR[OF measure_incseq, OF A]
     continuity_from_below[OF A] by simp
 
 lemma (in measure_space) measure_decseq:
@@ -264,10 +264,10 @@
   have A0: "0 \<le> \<mu> (A 0)" using A by auto
 
   have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))"
-    by (simp add: extreal_SUPR_uminus minus_extreal_def)
+    by (simp add: ereal_SUPR_uminus minus_ereal_def)
   also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))"
-    unfolding minus_extreal_def using A0 assms
-    by (subst SUPR_extreal_add) (auto simp add: measure_decseq)
+    unfolding minus_ereal_def using A0 assms
+    by (subst SUPR_ereal_add) (auto simp add: measure_decseq)
   also have "\<dots> = (SUP n. \<mu> (A 0 - A n))"
     using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto
   also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
@@ -280,7 +280,7 @@
   also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
     using A finite * by (simp, subst measure_Diff) auto
   finally show ?thesis
-    unfolding extreal_minus_eq_minus_iff using finite A0 by auto
+    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
 qed
 
 lemma (in measure_space) measure_insert:
@@ -489,7 +489,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"
-      using fin contr pos by (intro extreal_less_add) auto
+      using fin contr pos by (intro ereal_less_add) auto
     also have "\<dots> = \<mu> (T \<union> S)"
       using assms by (subst measure_additive) auto
     also have "\<dots> \<le> \<mu> (space M)"
@@ -1059,7 +1059,7 @@
   shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
   unfolding measure_additive[symmetric, OF measurable]
   using measurable(1,2)[THEN positive_measure]
-  using finite by (cases rule: extreal2_cases[of "\<mu> A" "\<mu> B"]) auto
+  using finite by (cases rule: ereal2_cases[of "\<mu> A" "\<mu> B"]) auto
 
 lemma (in measure_space) real_measure_finite_Union:
   assumes measurable:
@@ -1067,7 +1067,7 @@
   assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>"
   shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
   using finite measurable(2)[THEN positive_measure]
-  by (force intro!: setsum_real_of_extreal[symmetric]
+  by (force intro!: setsum_real_of_ereal[symmetric]
             simp: measure_setsum[OF measurable, symmetric])
 
 lemma (in measure_space) real_measure_Diff:
@@ -1088,7 +1088,7 @@
   shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
 proof -
   have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto
-  with summable_sums[OF summable_extreal_pos, of "\<lambda>i. \<mu> (A i)"]
+  with summable_sums[OF summable_ereal_pos, of "\<lambda>i. \<mu> (A i)"]
      measure_countably_additive[OF measurable]
   have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp
   moreover
@@ -1096,14 +1096,14 @@
     have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)"
       using measurable by (auto intro!: measure_mono)
     moreover have "0 \<le> \<mu> (A i)" using measurable by auto
-    ultimately have "\<mu> (A i) = extreal (real (\<mu> (A i)))"
+    ultimately have "\<mu> (A i) = ereal (real (\<mu> (A i)))"
       using finite by (cases "\<mu> (A i)") auto }
   moreover
   have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto
-  then have "\<mu> (\<Union>i. A i) = extreal (real (\<mu> (\<Union>i. A i)))"
+  then have "\<mu> (\<Union>i. A i) = ereal (real (\<mu> (\<Union>i. A i)))"
     using finite by (cases "\<mu> (\<Union>i. A i)") auto
   ultimately show ?thesis
-    unfolding sums_extreal[symmetric] by simp
+    unfolding sums_ereal[symmetric] by simp
 qed
 
 lemma (in measure_space) real_measure_subadditive:
@@ -1114,7 +1114,7 @@
   have "0 \<le> \<mu> (A \<union> B)" using measurable by auto
   then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
     using measure_subadditive[OF measurable] fin
-    by (cases rule: extreal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
+    by (cases rule: ereal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
 qed
 
 lemma (in measure_space) real_measure_setsum_singleton:
@@ -1123,24 +1123,24 @@
   shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
   using measure_finite_singleton[OF S] fin
   using positive_measure[OF S(2)]
-  by (force intro!: setsum_real_of_extreal[symmetric])
+  by (force intro!: setsum_real_of_ereal[symmetric])
 
 lemma (in measure_space) real_continuity_from_below:
   assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
 proof -
   have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
-  then have "extreal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
-    using fin by (auto intro: extreal_real')
+  then have "ereal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
+    using fin by (auto intro: ereal_real')
   then show ?thesis
     using continuity_from_below_Lim[OF A]
-    by (intro lim_real_of_extreal) simp
+    by (intro lim_real_of_ereal) simp
 qed
 
 lemma (in measure_space) continuity_from_above_Lim:
   assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)"
-  using LIMSEQ_extreal_INFI[OF measure_decseq, OF A]
+  using LIMSEQ_ereal_INFI[OF measure_decseq, OF A]
   using continuity_from_above[OF A fin] by simp
 
 lemma (in measure_space) real_continuity_from_above:
@@ -1151,11 +1151,11 @@
   moreover
   have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
     using A by (auto intro!: measure_mono)
-  ultimately have "extreal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
-    using fin by (auto intro: extreal_real')
+  ultimately have "ereal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
+    using fin by (auto intro: ereal_real')
   then show ?thesis
     using continuity_from_above_Lim[OF A fin]
-    by (intro lim_real_of_extreal) simp
+    by (intro lim_real_of_ereal) simp
 qed
 
 lemma (in measure_space) real_measure_countably_subadditive:
@@ -1167,11 +1167,11 @@
     moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto
     ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto }
   moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
-  ultimately have "extreal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. extreal (real (\<mu> (A i))))"
-    using measure_countably_subadditive[OF A] by (auto simp: extreal_real)
-  also have "\<dots> = extreal (\<Sum>i. real (\<mu> (A i)))"
+  ultimately have "ereal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. ereal (real (\<mu> (A i))))"
+    using measure_countably_subadditive[OF A] by (auto simp: ereal_real)
+  also have "\<dots> = ereal (\<Sum>i. real (\<mu> (A i)))"
     using A
-    by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin)
+    by (auto intro!: sums_unique[symmetric] sums_ereal[THEN iffD2] summable_sums summable_real_of_ereal fin)
   finally show ?thesis by simp
 qed
 
@@ -1199,14 +1199,14 @@
 definition (in finite_measure)
   "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)"
 
-lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = extreal (\<mu>' A)"
-  by (auto simp: \<mu>'_def extreal_real)
+lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = ereal (\<mu>' A)"
+  by (auto simp: \<mu>'_def ereal_real)
 
 lemma (in finite_measure) positive_measure'[simp, intro]: "0 \<le> \<mu>' A"
-  unfolding \<mu>'_def by (auto simp: real_of_extreal_pos)
+  unfolding \<mu>'_def by (auto simp: real_of_ereal_pos)
 
 lemma (in finite_measure) real_measure:
-  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = extreal r"
+  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = ereal r"
   using finite_measure[OF A] positive_measure[OF A] by (cases "\<mu> A") auto
 
 lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)"
@@ -1215,8 +1215,8 @@
   moreover then have "\<mu> A \<le> \<mu> (space M)"
     using sets_into_space by (auto intro!: measure_mono)
   ultimately show ?thesis
-    by (auto simp: \<mu>'_def intro!: real_of_extreal_positive_mono)
-qed (simp add: \<mu>'_def real_of_extreal_pos)
+    by (auto simp: \<mu>'_def intro!: real_of_ereal_positive_mono)
+qed (simp add: \<mu>'_def real_of_ereal_pos)
 
 lemma (in finite_measure) restricted_finite_measure:
   assumes "S \<in> sets M"
@@ -1302,8 +1302,8 @@
   note A[THEN subsetD, THEN finite_measure_eq, simp]
   note countable_UN[OF A, THEN finite_measure_eq, simp]
   from `summable (\<lambda>i. \<mu>' (A i))`
-  have "(\<lambda>i. extreal (\<mu>' (A i))) sums extreal (\<Sum>i. \<mu>' (A i))"
-    by (simp add: sums_extreal) (rule summable_sums)
+  have "(\<lambda>i. ereal (\<mu>' (A i))) sums ereal (\<Sum>i. \<mu>' (A i))"
+    by (simp add: sums_ereal) (rule summable_sums)
   from sums_unique[OF this, symmetric]
        measure_countably_subadditive[OF A]
   show ?thesis by simp
@@ -1440,26 +1440,26 @@
 qed
 
 lemma suminf_cmult_indicator:
-  fixes f :: "nat \<Rightarrow> extreal"
+  fixes f :: "nat \<Rightarrow> ereal"
   assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
   shows "(\<Sum>n. f n * indicator (A n) x) = f i"
 proof -
-  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)"
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
     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 :: extreal)"
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
     by (auto simp: setsum_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)"
-  proof (rule extreal_SUPI)
-    fix y :: extreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
+  proof (rule ereal_SUPI)
+    fix y :: ereal 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 (insert assms, simp)
   ultimately show ?thesis using assms
-    by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def)
+    by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
 qed
 
 lemma suminf_indicator:
   assumes "disjoint_family A"
-  shows "(\<Sum>n. indicator (A n) x :: extreal) = indicator (\<Union>i. A i) x"
+  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
 proof cases
   assume *: "x \<in> (\<Union>i. A i)"
   then obtain i where "x \<in> A i" by auto
--- a/src/HOL/Probability/Probability_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -54,7 +54,7 @@
   by (auto simp: distribution_def intro!: arg_cong[where f=prob])
 
 lemma (in prob_space) prob_space: "prob (space M) = 1"
-  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
+  using measure_space_1 unfolding \<mu>'_def by (simp add: one_ereal_def)
 
 lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
   using bounded_measure[of A] by (simp add: prob_space)
@@ -242,7 +242,7 @@
 
 lemma (in prob_space) distribution_prob_space:
   assumes X: "random_variable S X"
-  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
+  shows "prob_space (S\<lparr>measure := ereal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
 proof (rule prob_space_vimage)
   show "X \<in> measure_preserving M ?S"
     using X
@@ -252,10 +252,10 @@
 qed
 
 lemma (in prob_space) AE_distribution:
-  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
+  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := ereal \<circ> distribution X\<rparr>. Q x"
   shows "AE x. Q (X x)"
 proof -
-  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
+  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
   obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
     using assms unfolding X.almost_everywhere_def by auto
   from X[unfolded measurable_def] N show "AE x. Q (X x)"
@@ -410,9 +410,9 @@
 
 lemma (in prob_space) distribution_eq_translated_integral:
   assumes "random_variable S X" "A \<in> sets S"
-  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
+  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := ereal \<circ> distribution X\<rparr>) (indicator A)"
 proof -
-  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
+  interpret S: prob_space "S\<lparr>measure := ereal \<circ> distribution X\<rparr>"
     using assms(1) by (rule distribution_prob_space)
   show ?thesis
     using S.positive_integral_indicator(1)[of A] assms by simp
@@ -548,7 +548,7 @@
 
 lemma (in prob_space) joint_distribution_prob_space:
   assumes "random_variable MX X" "random_variable MY Y"
-  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
+  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
   using random_variable_pairI[OF assms] by (rule distribution_prob_space)
 
 
@@ -570,12 +570,12 @@
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
   shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
 proof -
-  have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
+  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
     using X by (intro finite_measure_eq[symmetric] in_P) auto
   also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
     using measure_times X by simp
-  also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
-    using X by (simp add: M.finite_measure_eq setprod_extreal)
+  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
+    using X by (simp add: M.finite_measure_eq setprod_ereal)
   finally show ?thesis by simp
 qed
 
@@ -610,9 +610,9 @@
 
 lemma (in prob_space) distribution_finite_prob_space:
   assumes "finite_random_variable MX X"
-  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
+  shows "finite_prob_space (MX\<lparr>measure := ereal \<circ> distribution X\<rparr>)"
 proof -
-  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
+  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>"
     using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
   interpret MX: finite_sigma_algebra MX
     using assms by auto
@@ -853,12 +853,12 @@
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)"
   shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
 proof -
-  have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
+  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
     using X by (intro finite_measure_eq[symmetric] in_P) auto
   also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
     using measure_finite_times X by simp
-  also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
-    using X by (simp add: M.finite_measure_eq setprod_extreal)
+  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
+    using X by (simp add: M.finite_measure_eq setprod_ereal)
   finally show ?thesis by simp
 qed
 
@@ -877,7 +877,7 @@
 lemma (in prob_space) joint_distribution_finite_prob_space:
   assumes X: "finite_random_variable MX X"
   assumes Y: "finite_random_variable MY Y"
-  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
+  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
   by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
 
 lemma finite_prob_space_eq:
@@ -1021,7 +1021,7 @@
       using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
     show "positive ?P (measure ?P)"
     proof (simp add: positive_def, safe)
-      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
+      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_ereal_def)
       fix B assume "B \<in> events"
       with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
       show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
@@ -1035,12 +1035,12 @@
         with `A \<in> events` have "F i \<in> events" by auto }
       moreover then have "range F \<subseteq> events" by auto
       moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
-        by (simp add: mult_commute divide_extreal_def)
+        by (simp add: mult_commute divide_ereal_def)
       moreover have "0 \<le> inverse (\<mu> A)"
         using real_measure[OF `A \<in> events`] by auto
       ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
         using measure_countably_additive[of F] F
-        by (auto simp: suminf_cmult_extreal)
+        by (auto simp: suminf_cmult_ereal)
     qed
   qed
 qed
@@ -1059,7 +1059,7 @@
 
 lemma (in finite_prob_space) finite_measure_space:
   fixes X :: "'a \<Rightarrow> 'x"
-  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
+  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X\<rparr>"
     (is "finite_measure_space ?S")
 proof (rule finite_measure_spaceI, simp_all)
   show "finite (X ` space M)" using finite_space by simp
@@ -1072,13 +1072,13 @@
 qed
 
 lemma (in finite_prob_space) finite_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
-  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
+  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X \<rparr>"
+  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def)
 
 lemma (in finite_prob_space) finite_product_measure_space:
   fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
   assumes "finite s1" "finite s2"
-  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
+  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = ereal \<circ> joint_distribution X Y\<rparr>"
     (is "finite_measure_space ?M")
 proof (rule finite_measure_spaceI, simp_all)
   show "finite (s1 \<times> s2)"
@@ -1094,14 +1094,14 @@
 lemma (in finite_prob_space) finite_product_measure_space_of_images:
   shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
                                 sets = Pow (X ` space M \<times> Y ` space M),
-                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
+                                measure = ereal \<circ> joint_distribution X Y \<rparr>"
   using finite_space by (auto intro!: finite_product_measure_space)
 
 lemma (in finite_prob_space) finite_product_prob_space_of_images:
   "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
-                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
+                       measure = ereal \<circ> joint_distribution X Y \<rparr>"
   (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def)
   have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
   thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
     by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
@@ -1129,7 +1129,7 @@
   by (simp add: pborel_def)
 
 interpretation pborel: prob_space pborel
-  by default (simp add: one_extreal_def pborel_def)
+  by default (simp add: one_ereal_def pborel_def)
 
 lemma pborel_prob: "pborel.prob A = (if A \<in> sets borel \<and> A \<subseteq> {0 ..< 1} then real (lborel.\<mu> A) else 0)"
   unfolding pborel.\<mu>'_def by (auto simp: pborel_def)
@@ -1171,11 +1171,11 @@
 subsection "Bernoulli space"
 
 definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV,
-  measure = extreal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
+  measure = ereal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
 
 interpretation bernoulli: finite_prob_space "bernoulli_space p" for p
   by (rule finite_prob_spaceI)
-     (auto simp: bernoulli_space_def UNIV_bool one_extreal_def setsum_Un_disjoint intro!: setsum_nonneg)
+     (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg)
 
 lemma bernoulli_measure:
   "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)"
--- a/src/HOL/Probability/Radon_Nikodym.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -23,7 +23,7 @@
     fix i have Ai: "A i \<in> sets M" using range by auto
     from measure positive_measure[OF this]
     show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
-      by (auto intro!: extreal_dense simp: extreal_0_gt_inverse)
+      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse)
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
     "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
@@ -40,15 +40,15 @@
       fix N
       have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)"
         using positive_measure[OF `A N \<in> sets M`] n[of N]
-        by (intro extreal_mult_right_mono) auto
+        by (intro ereal_mult_right_mono) auto
       also have "\<dots> \<le> (1 / 2) ^ Suc N"
         using measure[of N] n[of N]
-        by (cases rule: extreal2_cases[of "n N" "\<mu> (A N)"])
-           (simp_all add: inverse_eq_divide power_divide one_extreal_def extreal_power_divide)
+        by (cases rule: ereal2_cases[of "n N" "\<mu> (A N)"])
+           (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
       finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" .
       show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp
     qed
-    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_extreal by auto
+    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
   next
     { fix x assume "x \<in> space M"
       then obtain i where "x \<in> A i" using space[symmetric] by auto
@@ -65,14 +65,14 @@
     qed
   next
     show "?h \<in> borel_measurable M" using range n
-      by (auto intro!: borel_measurable_psuminf borel_measurable_extreal_times extreal_0_le_mult intro: less_imp_le)
+      by (auto intro!: borel_measurable_psuminf borel_measurable_ereal_times ereal_0_le_mult intro: less_imp_le)
   qed
 qed
 
 subsection "Absolutely continuous"
 
 definition (in measure_space)
-  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: extreal))"
+  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: ereal))"
 
 lemma (in measure_space) absolutely_continuous_AE:
   assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
@@ -199,7 +199,7 @@
       proof (induct n)
         case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
       next
-        case 0 with M'.empty_measure show ?case by (simp add: zero_extreal_def)
+        case 0 with M'.empty_measure show ?case by (simp add: zero_ereal_def)
       qed } note dA_less = this
     have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
     proof (rule incseq_SucI)
@@ -433,7 +433,7 @@
       by (intro positive_integral_suminf[symmetric]) auto
     also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)"
       using `\<And>x. 0 \<le> f x`
-      by (intro positive_integral_cong) (simp add: suminf_cmult_extreal suminf_indicator[OF `disjoint_family A`])
+      by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`])
     finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" .
     moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
       using M'.measure_countably_additive A by (simp add: comp_def)
@@ -447,14 +447,14 @@
       using A by (intro f_le_\<nu>) auto
     ultimately
     show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)"
-      by (subst suminf_extreal_minus) (simp_all add: positive_integral_positive)
+      by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive)
   next
     fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M"
-      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def extreal_le_minus_iff)
+      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff)
   next
     show "\<nu> (space M) - (\<integral>\<^isup>+ x. ?F (space M) x \<partial>M) \<noteq> \<infinity>" (is "?a - ?b \<noteq> _")
       using positive_integral_positive[of "?F (space M)"]
-      by (cases rule: extreal2_cases[of ?a ?b]) auto
+      by (cases rule: ereal2_cases[of ?a ?b]) auto
   qed
   then interpret M: finite_measure ?M
     where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
@@ -490,7 +490,7 @@
     def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
     ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
       using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"]
-      by (cases rule: extreal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
+      by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
          (simp_all add: field_simps)
     then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
     let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
@@ -501,7 +501,7 @@
         using `0 \<le> b` by (auto simp: positive_def)
       show "countably_additive ?Mb (measure ?Mb)"
         using `0 \<le> b` measure_countably_additive
-        by (auto simp: countably_additive_def suminf_cmult_extreal subset_eq)
+        by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq)
       show "measure ?Mb (space ?Mb) \<noteq> \<infinity>"
         using b by auto
     qed
@@ -513,7 +513,7 @@
     { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
       with *[OF this] have "b * \<mu> B \<le> ?t B"
         using M'.finite_measure b finite_measure M.positive_measure[OF B(1)]
-        by (cases rule: extreal2_cases[of "?t B" "b * \<mu> B"]) auto }
+        by (cases rule: ereal2_cases[of "?t B" "b * \<mu> B"]) auto }
     note bM_le_t = this
     let "?f0 x" = "f x + b * indicator A0 x"
     { fix A assume A: "A \<in> sets M"
@@ -543,8 +543,8 @@
         by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") auto
       finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
     hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
-      by (auto intro!: borel_measurable_indicator borel_measurable_extreal_add
-                       borel_measurable_extreal_times extreal_add_nonneg_nonneg)
+      by (auto intro!: borel_measurable_indicator borel_measurable_ereal_add
+                       borel_measurable_ereal_times ereal_add_nonneg_nonneg)
     have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>"
       "b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>"
       using `A0 \<in> sets M` b
@@ -552,12 +552,12 @@
         finite_measure_of_space M.finite_measure_of_space
       by auto
     have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
-      using M'.finite_measure_of_space pos_t unfolding extreal_less_minus_iff
+      using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff
       by (auto cong: positive_integral_cong)
     have  "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def
       using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
       using positive_integral_positive[of "?F (space M)"]
-      by (cases rule: extreal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
+      by (cases rule: ereal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
          (auto simp: field_simps mult_less_cancel_left)
     also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
       using space_less_A0 b
@@ -586,10 +586,10 @@
     then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def
       using `A0 \<in> sets M` by auto
     then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto
-    hence "0 < b * \<mu> A0" using b by (auto simp: extreal_zero_less_0_iff)
+    hence "0 < b * \<mu> A0" using b by (auto simp: ereal_zero_less_0_iff)
     with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
       using `f \<in> G`
-      by (intro extreal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive)
+      by (intro ereal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive)
     also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
       by (simp cong: positive_integral_cong)
     finally have "?y < integral\<^isup>P M ?f0" by simp
@@ -726,7 +726,7 @@
           using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
           using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"]
           using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono]
-          by (cases rule: extreal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
+          by (cases rule: ereal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
       qed }
     show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
     { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
@@ -764,7 +764,7 @@
     \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
   proof
     fix i
-    have indicator_eq: "\<And>f x A. (f x :: extreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
+    have indicator_eq: "\<And>f x A. (f x :: ereal) * 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))"
@@ -804,12 +804,12 @@
     have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
     have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
       "\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x"
-      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_extreal_times)
+      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
     have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
       using borel by (intro positive_integral_cong) (auto simp: indicator_def)
     also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)"
       using borel Qi Q0(1) `A \<in> sets M`
-      by (subst positive_integral_add) (auto simp del: extreal_infty_mult
+      by (subst positive_integral_add) (auto simp del: ereal_infty_mult
           simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
     also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)"
       by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto
@@ -866,7 +866,7 @@
   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_extreal_times)
+      using borel f_borel by (auto intro: borel_measurable_ereal_times)
     show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto
     fix A assume "A \<in> sets M"
     then show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
@@ -915,7 +915,7 @@
     finally have "AE x. f x \<le> g x"
       using pos borel positive_integral_PInf_AE[OF borel(2) g_fin]
       by (subst (asm) positive_integral_0_iff_AE)
-         (auto split: split_indicator simp: not_less extreal_minus_le_iff) }
+         (auto split: split_indicator simp: not_less ereal_minus_le_iff) }
   from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
   show "AE x. f x = g x" by auto
 qed
@@ -942,24 +942,24 @@
   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::extreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+  have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
     unfolding indicator_def by auto
   have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
     by (intro finite_density_unique[THEN iffD1] allI)
-       (auto intro!: borel_measurable_extreal_times f Int simp: *)
+       (auto intro!: borel_measurable_ereal_times f Int simp: *)
   moreover have "AE x. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
-    { fix f :: "'a \<Rightarrow> extreal" assume borel: "f \<in> borel_measurable M"
+    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
         and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
+      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
       have "(\<Union>i. ?A i) \<in> null_sets"
       proof (rule null_sets_UN)
-        fix i have "?A i \<in> sets M"
+        fix i ::nat have "?A i \<in> sets M"
           using borel Q0(1) by auto
-        have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. of_nat i * indicator (?A i) x \<partial>M)"
+        have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
           unfolding eq[OF `?A i \<in> sets M`]
           by (auto intro!: positive_integral_mono simp: indicator_def)
-        also have "\<dots> = of_nat i * \<mu> (?A i)"
+        also have "\<dots> = i * \<mu> (?A i)"
           using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
         also have "\<dots> < \<infinity>"
           using `?A i \<in> sets M`[THEN finite_measure] by auto
@@ -1067,7 +1067,7 @@
   interpret \<nu>: measure_space ?N
     where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
     and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
-  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
+  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M"
   { fix i j have "A i \<inter> Q j \<in> sets M"
     unfolding A_def using f Q
     apply (rule_tac Int)
@@ -1095,7 +1095,7 @@
         case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
       next
         case (real r)
-        with less_PInf_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by (auto simp: real_eq_of_nat)
+        with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat)
         then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"])
       next
         case MInf with x show ?thesis
@@ -1115,9 +1115,9 @@
     next
       case (Suc n)
       then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
-        (\<integral>\<^isup>+x. of_nat (Suc n) * indicator (Q j) x \<partial>M)"
+        (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
         by (auto intro!: positive_integral_mono simp: indicator_def A_def)
-      also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
+      also have "\<dots> = Suc n * \<mu> (Q j)"
         using Q by (auto intro!: positive_integral_cmult_indicator)
       also have "\<dots> < \<infinity>"
         using Q by (auto simp: real_eq_of_nat[symmetric])
@@ -1180,7 +1180,7 @@
   also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
     using RN_deriv(3)[OF \<nu>]
     by (auto intro!: positive_integral_cong_pos split: split_if_asm
-             simp: max_def extreal_mult_le_0_iff)
+             simp: max_def ereal_mult_le_0_iff)
   finally show ?thesis .
 qed
 
@@ -1287,18 +1287,18 @@
 proof -
   interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
   have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-  have minus_cong: "\<And>A B A' B'::extreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
+  have minus_cong: "\<And>A B A' B'::ereal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
   have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
   have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp
   { fix f :: "'a \<Rightarrow> real"
     { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>"
-      have "extreal (real (RN_deriv M \<nu> x)) * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
+      have "ereal (real (RN_deriv M \<nu> x)) * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
         by (simp add: mult_le_0_iff)
-      then have "RN_deriv M \<nu> x * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
-        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: extreal_real split: split_if_asm) }
-    then have "(\<integral>\<^isup>+x. extreal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * extreal (f x) \<partial>M)"
-              "(\<integral>\<^isup>+x. extreal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * extreal (- f x) \<partial>M)"
-      using RN_deriv_finite[OF \<nu>] unfolding extreal_mult_minus_right uminus_extreal.simps(1)[symmetric]
+      then have "RN_deriv M \<nu> x * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
+        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: ereal_real split: split_if_asm) }
+    then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (f x) \<partial>M)"
+              "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (- f x) \<partial>M)"
+      using RN_deriv_finite[OF \<nu>] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
       by (auto intro!: positive_integral_cong_AE) }
   note * = this
   show ?integral ?integrable
@@ -1311,7 +1311,7 @@
   assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>")
   assumes ac: "absolutely_continuous \<nu>"
   obtains D where "D \<in> borel_measurable M"
-    and "AE x. RN_deriv M \<nu> x = extreal (D x)"
+    and "AE x. RN_deriv M \<nu> x = ereal (D x)"
     and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x"
     and "\<And>x. 0 \<le> D x"
 proof
@@ -1342,13 +1342,13 @@
   qed
   ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>"
     using RN by (intro AE_iff_measurable[THEN iffD2]) auto
-  then show "AE x. RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
-    using RN(3) by (auto simp: extreal_real)
-  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
+  then show "AE x. RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
+    using RN(3) by (auto simp: ereal_real)
+  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
     using ac absolutely_continuous_AE[OF ms] by auto
 
   show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)"
-    using RN by (auto intro: real_of_extreal_pos)
+    using RN by (auto intro: real_of_ereal_pos)
 
   have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)"
     using RN by auto
@@ -1357,7 +1357,7 @@
   finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0"
     using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto
   with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)"
-    by (auto simp: zero_less_real_of_extreal le_less)
+    by (auto simp: zero_less_real_of_ereal le_less)
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_singleton:
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -8,7 +8,7 @@
   and not_empty[simp]: "S \<noteq> {}"
 
 definition (in finite_space) "M = \<lparr> space = S, sets = Pow S,
-  measure = \<lambda>A. extreal (card A / card S) \<rparr>"
+  measure = \<lambda>A. ereal (card A / card S) \<rparr>"
 
 lemma (in finite_space)
   shows space_M[simp]: "space M = S"
@@ -23,7 +23,7 @@
 qed (auto simp: M_def divide_nonneg_nonneg)
 
 sublocale finite_space \<subseteq> information_space M 2
-  by default (simp_all add: M_def one_extreal_def)
+  by default (simp_all add: M_def one_ereal_def)
 
 lemma (in finite_space) \<mu>'_eq[simp]: "\<mu>' A = (if A \<subseteq> S then card A / card S else 0)"
   unfolding \<mu>'_def by (auto simp: M_def)
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Jul 20 10:11:08 2011 +0200
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Jul 20 10:48:00 2011 +0200
@@ -201,13 +201,13 @@
 lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
    by (auto intro!: setsum_nonneg)
 
-sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
+sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
   by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset)
 
-sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
-  by default (simp add: one_extreal_def)
+sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
+  by default (simp add: one_ereal_def)
 
-sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>" b
+sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>" b
   by default simp
 
 lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> \<mu>' A = setsum p A"