merged
authorhoelzl
Wed, 20 Jul 2011 10:48:00 +0200
changeset 44796f651cb053927
parent 44789 6ca79a354c51
parent 44795 1165fe965da8
child 44797 3264fbfd87d6
merged
src/HOL/IsaMakefile
src/HOL/Library/Extended_Reals.thy
src/HOL/Library/Nat_Infinity.thy
     1.1 --- a/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Wed Jul 20 10:11:08 2011 +0200
     1.2 +++ b/src/HOL/HOLCF/FOCUS/Buffer_adm.thy	Wed Jul 20 10:48:00 2011 +0200
     1.3 @@ -8,7 +8,7 @@
     1.4  imports Buffer Stream_adm
     1.5  begin
     1.6  
     1.7 -declare Fin_0 [simp]
     1.8 +declare enat_0 [simp]
     1.9  
    1.10  lemma BufAC_Asm_d2: "a\<leadsto>s:BufAC_Asm ==> ? d. a=Md d"
    1.11  by (drule BufAC_Asm_unfold [THEN iffD1], auto)
    1.12 @@ -116,7 +116,7 @@
    1.13  done
    1.14  
    1.15  (*adm_BufAC_Asm,BufAC_Asm_antiton,adm_non_BufAC_Asm,BufAC_Asm_cong*)
    1.16 -lemma BufAC_Cmt_2stream_monoP: "f:BufEq ==> ? l. !i x s. s:BufAC_Asm --> x << s --> Fin (l i) < #x --> 
    1.17 +lemma BufAC_Cmt_2stream_monoP: "f:BufEq ==> ? l. !i x s. s:BufAC_Asm --> x << s --> enat (l i) < #x --> 
    1.18                       (x,f\<cdot>x):down_iterate BufAC_Cmt_F i --> 
    1.19                       (s,f\<cdot>s):down_iterate BufAC_Cmt_F i"
    1.20  apply (rule_tac x="%i. 2*i" in exI)
    1.21 @@ -139,10 +139,10 @@
    1.22         \<lbrakk>f \<in> BufEq;
    1.23            \<forall>x s. s \<in> BufAC_Asm \<longrightarrow>
    1.24                  x \<sqsubseteq> s \<longrightarrow>
    1.25 -                Fin (2 * i) < #x \<longrightarrow>
    1.26 +                enat (2 * i) < #x \<longrightarrow>
    1.27                  (x, f\<cdot>x) \<in> down_iterate BufAC_Cmt_F i \<longrightarrow>
    1.28                  (s, f\<cdot>s) \<in> down_iterate BufAC_Cmt_F i;
    1.29 -          Md d\<leadsto>\<bullet>\<leadsto>xa \<in> BufAC_Asm; Fin (2 * i) < #ya; f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>t;
    1.30 +          Md d\<leadsto>\<bullet>\<leadsto>xa \<in> BufAC_Asm; enat (2 * i) < #ya; f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>t;
    1.31            (ya, t) \<in> down_iterate BufAC_Cmt_F i; ya \<sqsubseteq> xa\<rbrakk>
    1.32         \<Longrightarrow> (xa, rt\<cdot>(f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>xa))) \<in> down_iterate BufAC_Cmt_F i
    1.33  *)
    1.34 @@ -158,11 +158,11 @@
    1.35  apply (erule subst)
    1.36  (*
    1.37   1. \<And>i d xa ya t ff ffa.
    1.38 -       \<lbrakk>f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>ffa\<cdot>ya; Fin (2 * i) < #ya;
    1.39 +       \<lbrakk>f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>ya) = d\<leadsto>ffa\<cdot>ya; enat (2 * i) < #ya;
    1.40            (ya, ffa\<cdot>ya) \<in> down_iterate BufAC_Cmt_F i; ya \<sqsubseteq> xa; f \<in> BufEq;
    1.41            \<forall>x s. s \<in> BufAC_Asm \<longrightarrow>
    1.42                  x \<sqsubseteq> s \<longrightarrow>
    1.43 -                Fin (2 * i) < #x \<longrightarrow>
    1.44 +                enat (2 * i) < #x \<longrightarrow>
    1.45                  (x, f\<cdot>x) \<in> down_iterate BufAC_Cmt_F i \<longrightarrow>
    1.46                  (s, f\<cdot>s) \<in> down_iterate BufAC_Cmt_F i;
    1.47            xa \<in> BufAC_Asm; ff \<in> BufEq; ffa \<in> BufEq\<rbrakk>
     2.1 --- a/src/HOL/HOLCF/FOCUS/Fstream.thy	Wed Jul 20 10:11:08 2011 +0200
     2.2 +++ b/src/HOL/HOLCF/FOCUS/Fstream.thy	Wed Jul 20 10:48:00 2011 +0200
     2.3 @@ -144,13 +144,13 @@
     2.4  by (simp add: fscons_def)
     2.5  
     2.6  lemma slen_fscons_eq:
     2.7 -        "(Fin (Suc n) < #x) = (? a y. x = a~> y & Fin n < #y)"
     2.8 +        "(enat (Suc n) < #x) = (? a y. x = a~> y & enat n < #y)"
     2.9  apply (simp add: fscons_def2 slen_scons_eq)
    2.10  apply (fast dest: not_Undef_is_Def [THEN iffD1] elim: DefE)
    2.11  done
    2.12  
    2.13  lemma slen_fscons_eq_rev:
    2.14 -        "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a~> y | #y < Fin (Suc n))"
    2.15 +        "(#x < enat (Suc (Suc n))) = (!a y. x ~= a~> y | #y < enat (Suc n))"
    2.16  apply (simp add: fscons_def2 slen_scons_eq_rev)
    2.17  apply (tactic {* step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1 *})
    2.18  apply (tactic {* step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1 *})
    2.19 @@ -163,7 +163,7 @@
    2.20  done
    2.21  
    2.22  lemma slen_fscons_less_eq:
    2.23 -        "(#(a~> y) < Fin (Suc (Suc n))) = (#y < Fin (Suc n))"
    2.24 +        "(#(a~> y) < enat (Suc (Suc n))) = (#y < enat (Suc n))"
    2.25  apply (subst slen_fscons_eq_rev)
    2.26  apply (fast dest!: fscons_inject [THEN iffD1])
    2.27  done
     3.1 --- a/src/HOL/HOLCF/FOCUS/Fstreams.thy	Wed Jul 20 10:11:08 2011 +0200
     3.2 +++ b/src/HOL/HOLCF/FOCUS/Fstreams.thy	Wed Jul 20 10:48:00 2011 +0200
     3.3 @@ -28,7 +28,7 @@
     3.4  
     3.5  definition
     3.6    jth           :: "nat => 'a fstream => 'a" where
     3.7 -  "jth = (%n s. if Fin n < #s then THE a. i_th n s = Def a else undefined)"
     3.8 +  "jth = (%n s. if enat n < #s then THE a. i_th n s = Def a else undefined)"
     3.9  
    3.10  definition
    3.11    first         :: "'a fstream => 'a" where
    3.12 @@ -36,7 +36,7 @@
    3.13  
    3.14  definition
    3.15    last          :: "'a fstream => 'a" where
    3.16 -  "last = (%s. case #s of Fin n => (if n~=0 then jth (THE k. Suc k = n) s else undefined))"
    3.17 +  "last = (%s. case #s of enat n => (if n~=0 then jth (THE k. Suc k = n) s else undefined))"
    3.18  
    3.19  
    3.20  abbreviation
    3.21 @@ -54,25 +54,25 @@
    3.22  lemma ft_fsingleton[simp]: "ft$(<a>) = Def a"
    3.23  by (simp add: fsingleton_def2)
    3.24  
    3.25 -lemma slen_fsingleton[simp]: "#(<a>) = Fin 1"
    3.26 -by (simp add: fsingleton_def2 inat_defs)
    3.27 +lemma slen_fsingleton[simp]: "#(<a>) = enat 1"
    3.28 +by (simp add: fsingleton_def2 enat_defs)
    3.29  
    3.30  lemma slen_fstreams[simp]: "#(<a> ooo s) = iSuc (#s)"
    3.31  by (simp add: fsingleton_def2)
    3.32  
    3.33  lemma slen_fstreams2[simp]: "#(s ooo <a>) = iSuc (#s)"
    3.34  apply (cases "#s")
    3.35 -apply (auto simp add: iSuc_Fin)
    3.36 +apply (auto simp add: iSuc_enat)
    3.37  apply (insert slen_sconc [of _ s "Suc 0" "<a>"], auto)
    3.38  by (simp add: sconc_def)
    3.39  
    3.40  lemma j_th_0_fsingleton[simp]:"jth 0 (<a>) = a"
    3.41  apply (simp add: fsingleton_def2 jth_def)
    3.42 -by (simp add: i_th_def Fin_0)
    3.43 +by (simp add: i_th_def enat_0)
    3.44  
    3.45  lemma jth_0[simp]: "jth 0 (<a> ooo s) = a"  
    3.46  apply (simp add: fsingleton_def2 jth_def)
    3.47 -by (simp add: i_th_def Fin_0)
    3.48 +by (simp add: i_th_def enat_0)
    3.49  
    3.50  lemma first_sconc[simp]: "first (<a> ooo s) = a"
    3.51  by (simp add: first_def)
    3.52 @@ -80,14 +80,14 @@
    3.53  lemma first_fsingleton[simp]: "first (<a>) = a"
    3.54  by (simp add: first_def)
    3.55  
    3.56 -lemma jth_n[simp]: "Fin n = #s ==> jth n (s ooo <a>) = a"
    3.57 +lemma jth_n[simp]: "enat n = #s ==> jth n (s ooo <a>) = a"
    3.58  apply (simp add: jth_def, auto)
    3.59  apply (simp add: i_th_def rt_sconc1)
    3.60 -by (simp add: inat_defs split: inat.splits)
    3.61 +by (simp add: enat_defs split: enat.splits)
    3.62  
    3.63 -lemma last_sconc[simp]: "Fin n = #s ==> last (s ooo <a>) = a"
    3.64 +lemma last_sconc[simp]: "enat n = #s ==> last (s ooo <a>) = a"
    3.65  apply (simp add: last_def)
    3.66 -apply (simp add: inat_defs split:inat.splits)
    3.67 +apply (simp add: enat_defs split:enat.splits)
    3.68  by (drule sym, auto)
    3.69  
    3.70  lemma last_fsingleton[simp]: "last (<a>) = a"
    3.71 @@ -97,18 +97,18 @@
    3.72  by (simp add: first_def jth_def)
    3.73  
    3.74  lemma last_UU[simp]:"last UU = undefined"
    3.75 -by (simp add: last_def jth_def inat_defs)
    3.76 +by (simp add: last_def jth_def enat_defs)
    3.77  
    3.78 -lemma last_infinite[simp]:"#s = Infty ==> last s = undefined"
    3.79 +lemma last_infinite[simp]:"#s = \<infinity> ==> last s = undefined"
    3.80  by (simp add: last_def)
    3.81  
    3.82 -lemma jth_slen_lemma1:"n <= k & Fin n = #s ==> jth k s = undefined"
    3.83 -by (simp add: jth_def inat_defs split:inat.splits, auto)
    3.84 +lemma jth_slen_lemma1:"n <= k & enat n = #s ==> jth k s = undefined"
    3.85 +by (simp add: jth_def enat_defs split:enat.splits, auto)
    3.86  
    3.87  lemma jth_UU[simp]:"jth n UU = undefined" 
    3.88  by (simp add: jth_def)
    3.89  
    3.90 -lemma ext_last:"[|s ~= UU; Fin (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" 
    3.91 +lemma ext_last:"[|s ~= UU; enat (Suc n) = #s|] ==> (stream_take n$s) ooo <(last s)> = s" 
    3.92  apply (simp add: last_def)
    3.93  apply (case_tac "#s", auto)
    3.94  apply (simp add: fsingleton_def2)
     4.1 --- a/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Wed Jul 20 10:11:08 2011 +0200
     4.2 +++ b/src/HOL/HOLCF/FOCUS/Stream_adm.thy	Wed Jul 20 10:48:00 2011 +0200
     4.3 @@ -10,14 +10,14 @@
     4.4  
     4.5  definition
     4.6    stream_monoP  :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
     4.7 -  "stream_monoP F = (\<exists>Q i. \<forall>P s. Fin i \<le> #s \<longrightarrow>
     4.8 +  "stream_monoP F = (\<exists>Q i. \<forall>P s. enat i \<le> #s \<longrightarrow>
     4.9                      (s \<in> F P) = (stream_take i\<cdot>s \<in> Q \<and> iterate i\<cdot>rt\<cdot>s \<in> P))"
    4.10  
    4.11  definition
    4.12    stream_antiP  :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
    4.13    "stream_antiP F = (\<forall>P x. \<exists>Q i.
    4.14 -                (#x  < Fin i \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow> y \<in> F P \<longrightarrow> x \<in> F P)) \<and>
    4.15 -                (Fin i <= #x \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow>
    4.16 +                (#x  < enat i \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow> y \<in> F P \<longrightarrow> x \<in> F P)) \<and>
    4.17 +                (enat i <= #x \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow>
    4.18                  (y \<in> F P) = (stream_take i\<cdot>y \<in> Q \<and> iterate i\<cdot>rt\<cdot>y \<in> P))))"
    4.19  
    4.20  definition
    4.21 @@ -57,7 +57,7 @@
    4.22  lemma flatstream_adm_lemma:
    4.23    assumes 1: "Porder.chain Y"
    4.24    assumes 2: "!i. P (Y i)"
    4.25 -  assumes 3: "(!!Y. [| Porder.chain Y; !i. P (Y i); !k. ? j. Fin k < #((Y j)::'a::flat stream)|]
    4.26 +  assumes 3: "(!!Y. [| Porder.chain Y; !i. P (Y i); !k. ? j. enat k < #((Y j)::'a::flat stream)|]
    4.27    ==> P(LUB i. Y i))"
    4.28    shows "P(LUB i. Y i)"
    4.29  apply (rule increasing_chain_adm_lemma [of _ P, OF 1 2])
    4.30 @@ -74,12 +74,12 @@
    4.31  apply (   erule spec)
    4.32  apply (  assumption)
    4.33  apply ( assumption)
    4.34 -apply (metis inat_ord_code(4) slen_infinite)
    4.35 +apply (metis enat_ord_code(4) slen_infinite)
    4.36  done
    4.37  
    4.38  (* should be without reference to stream length? *)
    4.39  lemma flatstream_admI: "[|(!!Y. [| Porder.chain Y; !i. P (Y i); 
    4.40 - !k. ? j. Fin k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P"
    4.41 + !k. ? j. enat k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P"
    4.42  apply (unfold adm_def)
    4.43  apply (intro strip)
    4.44  apply (erule (1) flatstream_adm_lemma)
    4.45 @@ -87,13 +87,13 @@
    4.46  done
    4.47  
    4.48  
    4.49 -(* context (theory "Nat_InFinity");*)
    4.50 -lemma ile_lemma: "Fin (i + j) <= x ==> Fin i <= x"
    4.51 +(* context (theory "Extended_Nat");*)
    4.52 +lemma ile_lemma: "enat (i + j) <= x ==> enat i <= x"
    4.53    by (rule order_trans) auto
    4.54  
    4.55  lemma stream_monoP2I:
    4.56  "!!X. stream_monoP F ==> !i. ? l. !x y. 
    4.57 -  Fin l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i"
    4.58 +  enat l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i"
    4.59  apply (unfold stream_monoP_def)
    4.60  apply (safe)
    4.61  apply (rule_tac x="i*ia" in exI)
    4.62 @@ -120,7 +120,7 @@
    4.63  done
    4.64  
    4.65  lemma stream_monoP2_gfp_admI: "[| !i. ? l. !x y. 
    4.66 - Fin l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i;
    4.67 + enat l <= #x --> (x::'a::flat stream) << y --> x:down_iterate F i --> y:down_iterate F i;
    4.68      down_cont F |] ==> adm (%x. x:gfp F)"
    4.69  apply (erule INTER_down_iterate_is_gfp [THEN ssubst]) (* cont *)
    4.70  apply (simp (no_asm))
    4.71 @@ -153,7 +153,7 @@
    4.72  apply (intro strip)
    4.73  apply (erule allE, erule all_dupE, erule exE, erule exE)
    4.74  apply (erule conjE)
    4.75 -apply (case_tac "#x < Fin i")
    4.76 +apply (case_tac "#x < enat i")
    4.77  apply ( fast)
    4.78  apply (unfold linorder_not_less)
    4.79  apply (drule (1) mp)
     5.1 --- a/src/HOL/HOLCF/IsaMakefile	Wed Jul 20 10:11:08 2011 +0200
     5.2 +++ b/src/HOL/HOLCF/IsaMakefile	Wed Jul 20 10:48:00 2011 +0200
     5.3 @@ -134,7 +134,7 @@
     5.4  HOLCF-ex: HOLCF $(LOG)/HOLCF-ex.gz
     5.5  
     5.6  $(LOG)/HOLCF-ex.gz: $(OUT)/HOLCF \
     5.7 -  ../Library/Nat_Infinity.thy \
     5.8 +  ../Library/Extended_Nat.thy \
     5.9    ex/Concurrency_Monad.thy \
    5.10    ex/Dagstuhl.thy \
    5.11    ex/Dnat.thy \
     6.1 --- a/src/HOL/HOLCF/Library/Stream.thy	Wed Jul 20 10:11:08 2011 +0200
     6.2 +++ b/src/HOL/HOLCF/Library/Stream.thy	Wed Jul 20 10:48:00 2011 +0200
     6.3 @@ -5,7 +5,7 @@
     6.4  header {* General Stream domain *}
     6.5  
     6.6  theory Stream
     6.7 -imports HOLCF "~~/src/HOL/Library/Nat_Infinity"
     6.8 +imports HOLCF "~~/src/HOL/Library/Extended_Nat"
     6.9  begin
    6.10  
    6.11  default_sort pcpo
    6.12 @@ -22,8 +22,8 @@
    6.13                                       If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)"
    6.14  
    6.15  definition
    6.16 -  slen :: "'a stream \<Rightarrow> inat"  ("#_" [1000] 1000) where
    6.17 -  "#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
    6.18 +  slen :: "'a stream \<Rightarrow> enat"  ("#_" [1000] 1000) where
    6.19 +  "#s = (if stream_finite s then enat (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
    6.20  
    6.21  
    6.22  (* concatenation *)
    6.23 @@ -39,7 +39,7 @@
    6.24  definition
    6.25    sconc :: "'a stream => 'a stream => 'a stream"  (infixr "ooo" 65) where
    6.26    "s1 ooo s2 = (case #s1 of
    6.27 -                  Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
    6.28 +                  enat n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
    6.29                 | \<infinity>     \<Rightarrow> s1)"
    6.30  
    6.31  primrec constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream"
    6.32 @@ -51,7 +51,7 @@
    6.33  definition
    6.34    constr_sconc  :: "'a stream => 'a stream => 'a stream" where (* constructive *)
    6.35    "constr_sconc s1 s2 = (case #s1 of
    6.36 -                          Fin n \<Rightarrow> constr_sconc' n s1 s2
    6.37 +                          enat n \<Rightarrow> constr_sconc' n s1 s2
    6.38                          | \<infinity>    \<Rightarrow> s1)"
    6.39  
    6.40  
    6.41 @@ -327,12 +327,12 @@
    6.42  section "slen"
    6.43  
    6.44  lemma slen_empty [simp]: "#\<bottom> = 0"
    6.45 -by (simp add: slen_def stream.finite_def zero_inat_def Least_equality)
    6.46 +by (simp add: slen_def stream.finite_def zero_enat_def Least_equality)
    6.47  
    6.48  lemma slen_scons [simp]: "x ~= \<bottom> ==> #(x&&xs) = iSuc (#xs)"
    6.49  apply (case_tac "stream_finite (x && xs)")
    6.50  apply (simp add: slen_def, auto)
    6.51 -apply (simp add: stream.finite_def, auto simp add: iSuc_Fin)
    6.52 +apply (simp add: stream.finite_def, auto simp add: iSuc_enat)
    6.53  apply (rule Least_Suc2, auto)
    6.54  (*apply (drule sym)*)
    6.55  (*apply (drule sym scons_eq_UU [THEN iffD1],simp)*)
    6.56 @@ -340,13 +340,13 @@
    6.57  apply (simp add: slen_def, auto)
    6.58  by (drule stream_finite_lemma1,auto)
    6.59  
    6.60 -lemma slen_less_1_eq: "(#x < Fin (Suc 0)) = (x = \<bottom>)"
    6.61 -by (cases x, auto simp add: Fin_0 iSuc_Fin[THEN sym])
    6.62 +lemma slen_less_1_eq: "(#x < enat (Suc 0)) = (x = \<bottom>)"
    6.63 +by (cases x, auto simp add: enat_0 iSuc_enat[THEN sym])
    6.64  
    6.65  lemma slen_empty_eq: "(#x = 0) = (x = \<bottom>)"
    6.66  by (cases x, auto)
    6.67  
    6.68 -lemma slen_scons_eq: "(Fin (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  Fin n < #y)"
    6.69 +lemma slen_scons_eq: "(enat (Suc n) < #x) = (? a y. x = a && y &  a ~= \<bottom> &  enat n < #y)"
    6.70  apply (auto, case_tac "x=UU",auto)
    6.71  apply (drule stream_exhaust_eq [THEN iffD1], auto)
    6.72  apply (case_tac "#y") apply simp_all
    6.73 @@ -359,18 +359,18 @@
    6.74  lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \<infinity>"
    6.75  by (simp add: slen_def)
    6.76  
    6.77 -lemma slen_scons_eq_rev: "(#x < Fin (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < Fin (Suc n))"
    6.78 +lemma slen_scons_eq_rev: "(#x < enat (Suc (Suc n))) = (!a y. x ~= a && y |  a = \<bottom> |  #y < enat (Suc n))"
    6.79   apply (cases x, auto)
    6.80 -   apply (simp add: zero_inat_def)
    6.81 -  apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
    6.82 - apply (case_tac "#stream") apply (simp_all add: iSuc_Fin)
    6.83 +   apply (simp add: zero_enat_def)
    6.84 +  apply (case_tac "#stream") apply (simp_all add: iSuc_enat)
    6.85 + apply (case_tac "#stream") apply (simp_all add: iSuc_enat)
    6.86  done
    6.87  
    6.88  lemma slen_take_lemma4 [rule_format]:
    6.89 -  "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n"
    6.90 -apply (induct n, auto simp add: Fin_0)
    6.91 +  "!s. stream_take n$s ~= s --> #(stream_take n$s) = enat n"
    6.92 +apply (induct n, auto simp add: enat_0)
    6.93  apply (case_tac "s=UU", simp)
    6.94 -by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_Fin)
    6.95 +by (drule stream_exhaust_eq [THEN iffD1], auto simp add: iSuc_enat)
    6.96  
    6.97  (*
    6.98  lemma stream_take_idempotent [simp]:
    6.99 @@ -390,41 +390,41 @@
   6.100  by (simp add: chain_def,simp)
   6.101  *)
   6.102  
   6.103 -lemma slen_take_eq: "ALL x. (Fin n < #x) = (stream_take n\<cdot>x ~= x)"
   6.104 +lemma slen_take_eq: "ALL x. (enat n < #x) = (stream_take n\<cdot>x ~= x)"
   6.105  apply (induct_tac n, auto)
   6.106 -apply (simp add: Fin_0, clarsimp)
   6.107 +apply (simp add: enat_0, clarsimp)
   6.108  apply (drule not_sym)
   6.109  apply (drule slen_empty_eq [THEN iffD1], simp)
   6.110  apply (case_tac "x=UU", simp)
   6.111  apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
   6.112  apply (erule_tac x="y" in allE, auto)
   6.113 -apply (simp_all add: not_less iSuc_Fin)
   6.114 +apply (simp_all add: not_less iSuc_enat)
   6.115  apply (case_tac "#y") apply simp_all
   6.116  apply (case_tac "x=UU", simp)
   6.117  apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)
   6.118  apply (erule_tac x="y" in allE, simp)
   6.119  apply (case_tac "#y") by simp_all
   6.120  
   6.121 -lemma slen_take_eq_rev: "(#x <= Fin n) = (stream_take n\<cdot>x = x)"
   6.122 +lemma slen_take_eq_rev: "(#x <= enat n) = (stream_take n\<cdot>x = x)"
   6.123  by (simp add: linorder_not_less [symmetric] slen_take_eq)
   6.124  
   6.125 -lemma slen_take_lemma1: "#x = Fin n ==> stream_take n\<cdot>x = x"
   6.126 +lemma slen_take_lemma1: "#x = enat n ==> stream_take n\<cdot>x = x"
   6.127  by (rule slen_take_eq_rev [THEN iffD1], auto)
   6.128  
   6.129  lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)"
   6.130  apply (cases s1)
   6.131   by (cases s2, simp+)+
   6.132  
   6.133 -lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"
   6.134 +lemma slen_take_lemma5: "#(stream_take n$s) <= enat n"
   6.135  apply (case_tac "stream_take n$s = s")
   6.136   apply (simp add: slen_take_eq_rev)
   6.137  by (simp add: slen_take_lemma4)
   6.138  
   6.139 -lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = Fin i"
   6.140 +lemma slen_take_lemma2: "!x. ~stream_finite x --> #(stream_take i\<cdot>x) = enat i"
   6.141  apply (simp add: stream.finite_def, auto)
   6.142  by (simp add: slen_take_lemma4)
   6.143  
   6.144 -lemma slen_infinite: "stream_finite x = (#x ~= Infty)"
   6.145 +lemma slen_infinite: "stream_finite x = (#x ~= \<infinity>)"
   6.146  by (simp add: slen_def)
   6.147  
   6.148  lemma slen_mono_lemma: "stream_finite s ==> ALL t. s << t --> #s <= #t"
   6.149 @@ -443,19 +443,19 @@
   6.150  lemma iterate_lemma: "F$(iterate n$F$x) = iterate n$F$(F$x)"
   6.151  by (insert iterate_Suc2 [of n F x], auto)
   6.152  
   6.153 -lemma slen_rt_mult [rule_format]: "!x. Fin (i + j) <= #x --> Fin j <= #(iterate i$rt$x)"
   6.154 +lemma slen_rt_mult [rule_format]: "!x. enat (i + j) <= #x --> enat j <= #(iterate i$rt$x)"
   6.155  apply (induct i, auto)
   6.156 -apply (case_tac "x=UU", auto simp add: zero_inat_def)
   6.157 +apply (case_tac "x=UU", auto simp add: zero_enat_def)
   6.158  apply (drule stream_exhaust_eq [THEN iffD1], auto)
   6.159  apply (erule_tac x="y" in allE, auto)
   6.160 -apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_Fin)
   6.161 +apply (simp add: not_le) apply (case_tac "#y") apply (simp_all add: iSuc_enat)
   6.162  by (simp add: iterate_lemma)
   6.163  
   6.164  lemma slen_take_lemma3 [rule_format]:
   6.165 -  "!(x::'a::flat stream) y. Fin n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
   6.166 +  "!(x::'a::flat stream) y. enat n <= #x --> x << y --> stream_take n\<cdot>x = stream_take n\<cdot>y"
   6.167  apply (induct_tac n, auto)
   6.168  apply (case_tac "x=UU", auto)
   6.169 -apply (simp add: zero_inat_def)
   6.170 +apply (simp add: zero_enat_def)
   6.171  apply (simp add: Suc_ile_eq)
   6.172  apply (case_tac "y=UU", clarsimp)
   6.173  apply (drule stream_exhaust_eq [THEN iffD1], clarsimp)+
   6.174 @@ -478,7 +478,7 @@
   6.175  apply (subgoal_tac "stream_take n$s ~=s")
   6.176   apply (insert slen_take_lemma4 [of n s],auto)
   6.177  apply (cases s, simp)
   6.178 -by (simp add: slen_take_lemma4 iSuc_Fin)
   6.179 +by (simp add: slen_take_lemma4 iSuc_enat)
   6.180  
   6.181  (* ----------------------------------------------------------------------- *)
   6.182  (* theorems about smap                                                     *)
   6.183 @@ -546,7 +546,7 @@
   6.184  lemma i_rt_mono: "x << s ==> i_rt n x  << i_rt n s"
   6.185  by (simp add: i_rt_def monofun_rt_mult)
   6.186  
   6.187 -lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)"
   6.188 +lemma i_rt_ij_lemma: "enat (i + j) <= #x ==> enat j <= #(i_rt i x)"
   6.189  by (simp add: i_rt_def slen_rt_mult)
   6.190  
   6.191  lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)"
   6.192 @@ -566,14 +566,14 @@
   6.193   apply (subgoal_tac "#(i_rt n s)=0")
   6.194    apply (case_tac "stream_take n$s = s",simp+)
   6.195    apply (insert slen_take_eq [rule_format,of n s],simp)
   6.196 -  apply (cases "#s") apply (simp_all add: zero_inat_def)
   6.197 +  apply (cases "#s") apply (simp_all add: zero_enat_def)
   6.198    apply (simp add: slen_take_eq)
   6.199    apply (cases "#s")
   6.200    using i_rt_take_lemma1 [of n s]
   6.201 -  apply (simp_all add: zero_inat_def)
   6.202 +  apply (simp_all add: zero_enat_def)
   6.203    done
   6.204  
   6.205 -lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
   6.206 +lemma i_rt_lemma_slen: "#s=enat n ==> i_rt n s = UU"
   6.207  by (simp add: i_rt_slen slen_take_lemma1)
   6.208  
   6.209  lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s"
   6.210 @@ -581,29 +581,29 @@
   6.211   apply (cases s, auto simp del: i_rt_Suc)
   6.212  by (simp add: i_rt_Suc_back stream_finite_rt_eq)+
   6.213  
   6.214 -lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl &
   6.215 -                            #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j
   6.216 -                                              --> Fin (j + t) = #x"
   6.217 +lemma take_i_rt_len_lemma: "ALL sl x j t. enat sl = #x & n <= sl &
   6.218 +                            #(stream_take n$x) = enat t & #(i_rt n x)= enat j
   6.219 +                                              --> enat (j + t) = #x"
   6.220  apply (induct n, auto)
   6.221 - apply (simp add: zero_inat_def)
   6.222 + apply (simp add: zero_enat_def)
   6.223  apply (case_tac "x=UU",auto)
   6.224 - apply (simp add: zero_inat_def)
   6.225 + apply (simp add: zero_enat_def)
   6.226  apply (drule stream_exhaust_eq [THEN iffD1],clarsimp)
   6.227 -apply (subgoal_tac "EX k. Fin k = #y",clarify)
   6.228 +apply (subgoal_tac "EX k. enat k = #y",clarify)
   6.229   apply (erule_tac x="k" in allE)
   6.230   apply (erule_tac x="y" in allE,auto)
   6.231   apply (erule_tac x="THE p. Suc p = t" in allE,auto)
   6.232 -   apply (simp add: iSuc_def split: inat.splits)
   6.233 -  apply (simp add: iSuc_def split: inat.splits)
   6.234 +   apply (simp add: iSuc_def split: enat.splits)
   6.235 +  apply (simp add: iSuc_def split: enat.splits)
   6.236    apply (simp only: the_equality)
   6.237 - apply (simp add: iSuc_def split: inat.splits)
   6.238 + apply (simp add: iSuc_def split: enat.splits)
   6.239   apply force
   6.240 -apply (simp add: iSuc_def split: inat.splits)
   6.241 +apply (simp add: iSuc_def split: enat.splits)
   6.242  done
   6.243  
   6.244  lemma take_i_rt_len:
   6.245 -"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==>
   6.246 -    Fin (j + t) = #x"
   6.247 +"[| enat sl = #x; n <= sl; #(stream_take n$x) = enat t; #(i_rt n x) = enat j |] ==>
   6.248 +    enat (j + t) = #x"
   6.249  by (blast intro: take_i_rt_len_lemma [rule_format])
   6.250  
   6.251  
   6.252 @@ -690,13 +690,13 @@
   6.253  (* ----------------------------------------------------------------------- *)
   6.254  
   6.255  lemma UU_sconc [simp]: " UU ooo s = s "
   6.256 -by (simp add: sconc_def zero_inat_def)
   6.257 +by (simp add: sconc_def zero_enat_def)
   6.258  
   6.259  lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
   6.260  by auto
   6.261  
   6.262  lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
   6.263 -apply (simp add: sconc_def zero_inat_def iSuc_def split: inat.splits, auto)
   6.264 +apply (simp add: sconc_def zero_enat_def iSuc_def split: enat.splits, auto)
   6.265  apply (rule someI2_ex,auto)
   6.266   apply (rule_tac x="x && y" in exI,auto)
   6.267  apply (simp add: i_rt_Suc_forw)
   6.268 @@ -704,21 +704,21 @@
   6.269  by (drule stream_exhaust_eq [THEN iffD1],auto)
   6.270  
   6.271  lemma ex_sconc [rule_format]:
   6.272 -  "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)"
   6.273 +  "ALL k y. #x = enat k --> (EX w. stream_take k$w = x & i_rt k w = y)"
   6.274  apply (case_tac "#x")
   6.275   apply (rule stream_finite_ind [of x],auto)
   6.276    apply (simp add: stream.finite_def)
   6.277    apply (drule slen_take_lemma1,blast)
   6.278 - apply (simp_all add: zero_inat_def iSuc_def split: inat.splits)
   6.279 + apply (simp_all add: zero_enat_def iSuc_def split: enat.splits)
   6.280  apply (erule_tac x="y" in allE,auto)
   6.281  by (rule_tac x="a && w" in exI,auto)
   6.282  
   6.283 -lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"
   6.284 -apply (simp add: sconc_def split: inat.splits, arith?,auto)
   6.285 +lemma rt_sconc1: "enat n = #x ==> i_rt n (x ooo y) = y"
   6.286 +apply (simp add: sconc_def split: enat.splits, arith?,auto)
   6.287  apply (rule someI2_ex,auto)
   6.288  by (drule ex_sconc,simp)
   6.289  
   6.290 -lemma sconc_inj2: "\<lbrakk>Fin n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
   6.291 +lemma sconc_inj2: "\<lbrakk>enat n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
   6.292  apply (frule_tac y=y in rt_sconc1)
   6.293  by (auto elim: rt_sconc1)
   6.294  
   6.295 @@ -734,7 +734,7 @@
   6.296   apply (simp add: i_rt_slen)
   6.297  by (simp add: sconc_def)
   6.298  
   6.299 -lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x"
   6.300 +lemma stream_take_sconc [simp]: "enat n = #x ==> stream_take n$(x ooo y) = x"
   6.301  apply (simp add: sconc_def)
   6.302  apply (cases "#x")
   6.303  apply auto
   6.304 @@ -743,7 +743,7 @@
   6.305  
   6.306  lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y"
   6.307  apply (cases "#x",auto)
   6.308 - apply (simp add: sconc_def iSuc_Fin)
   6.309 + apply (simp add: sconc_def iSuc_enat)
   6.310   apply (rule someI2_ex)
   6.311    apply (drule ex_sconc, simp)
   6.312   apply (rule someI2_ex, auto)
   6.313 @@ -799,9 +799,9 @@
   6.314  by (cases s, auto)
   6.315  
   6.316  lemma i_th_sconc_lemma [rule_format]:
   6.317 -  "ALL x y. Fin n < #x --> i_th n (x ooo y) = i_th n x"
   6.318 +  "ALL x y. enat n < #x --> i_th n (x ooo y) = i_th n x"
   6.319  apply (induct_tac n, auto)
   6.320 -apply (simp add: Fin_0 i_th_def)
   6.321 +apply (simp add: enat_0 i_th_def)
   6.322  apply (simp add: slen_empty_eq ft_sconc)
   6.323  apply (simp add: i_th_def)
   6.324  apply (case_tac "x=UU",auto)
   6.325 @@ -849,16 +849,16 @@
   6.326  (* ----------------------------------------------------------------------- *)
   6.327  
   6.328  lemma slen_sconc_finite1:
   6.329 -  "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty"
   6.330 -apply (case_tac "#y ~= Infty",auto)
   6.331 +  "[| #(x ooo y) = \<infinity>; enat n = #x |] ==> #y = \<infinity>"
   6.332 +apply (case_tac "#y ~= \<infinity>",auto)
   6.333  apply (drule_tac y=y in rt_sconc1)
   6.334  apply (insert stream_finite_i_rt [of n "x ooo y"])
   6.335  by (simp add: slen_infinite)
   6.336  
   6.337 -lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty"
   6.338 +lemma slen_sconc_infinite1: "#x=\<infinity> ==> #(x ooo y) = \<infinity>"
   6.339  by (simp add: sconc_def)
   6.340  
   6.341 -lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty"
   6.342 +lemma slen_sconc_infinite2: "#y=\<infinity> ==> #(x ooo y) = \<infinity>"
   6.343  apply (case_tac "#x")
   6.344   apply (simp add: sconc_def)
   6.345   apply (rule someI2_ex)
   6.346 @@ -868,7 +868,7 @@
   6.347   apply (fastsimp simp add: slen_infinite,auto)
   6.348  by (simp add: sconc_def)
   6.349  
   6.350 -lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)"
   6.351 +lemma sconc_finite: "(#x~=\<infinity> & #y~=\<infinity>) = (#(x ooo y)~=\<infinity>)"
   6.352  apply auto
   6.353    apply (metis not_Infty_eq slen_sconc_finite1)
   6.354   apply (metis not_Infty_eq slen_sconc_infinite1)
   6.355 @@ -877,7 +877,7 @@
   6.356  
   6.357  (* ----------------------------------------------------------------------- *)
   6.358  
   6.359 -lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k"
   6.360 +lemma slen_sconc_mono3: "[| enat n = #x; enat k = #(x ooo y) |] ==> n <= k"
   6.361  apply (insert slen_mono [of "x" "x ooo y"])
   6.362  apply (cases "#x") apply simp_all
   6.363  apply (cases "#(x ooo y)") apply simp_all
   6.364 @@ -887,10 +887,10 @@
   6.365     subsection "finite slen"
   6.366  (* ----------------------------------------------------------------------- *)
   6.367  
   6.368 -lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)"
   6.369 +lemma slen_sconc: "[| enat n = #x; enat m = #y |] ==> #(x ooo y) = enat (n + m)"
   6.370  apply (case_tac "#(x ooo y)")
   6.371   apply (frule_tac y=y in rt_sconc1)
   6.372 - apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp)
   6.373 + apply (insert take_i_rt_len [of "THE j. enat j = #(x ooo y)" "x ooo y" n n m],simp)
   6.374   apply (insert slen_sconc_mono3 [of n x _ y],simp)
   6.375  by (insert sconc_finite [of x y],auto)
   6.376  
   6.377 @@ -934,7 +934,7 @@
   6.378  
   6.379  lemma contlub_sconc_lemma:
   6.380    "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
   6.381 -apply (case_tac "#x=Infty")
   6.382 +apply (case_tac "#x=\<infinity>")
   6.383   apply (simp add: sconc_def)
   6.384  apply (drule finite_lub_sconc,auto simp add: slen_infinite)
   6.385  done
   6.386 @@ -948,7 +948,7 @@
   6.387  (* ----------------------------------------------------------------------- *)
   6.388  
   6.389  lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
   6.390 -by (simp add: constr_sconc_def zero_inat_def)
   6.391 +by (simp add: constr_sconc_def zero_enat_def)
   6.392  
   6.393  lemma "x ooo y = constr_sconc x y"
   6.394  apply (case_tac "#x")
   6.395 @@ -956,7 +956,7 @@
   6.396    defer 1
   6.397    apply (simp add: constr_sconc_def del: scons_sconc)
   6.398    apply (case_tac "#s")
   6.399 -   apply (simp add: iSuc_Fin)
   6.400 +   apply (simp add: iSuc_enat)
   6.401     apply (case_tac "a=UU",auto simp del: scons_sconc)
   6.402     apply (simp)
   6.403    apply (simp add: sconc_def)
     7.1 --- a/src/HOL/IsaMakefile	Wed Jul 20 10:11:08 2011 +0200
     7.2 +++ b/src/HOL/IsaMakefile	Wed Jul 20 10:48:00 2011 +0200
     7.3 @@ -444,25 +444,25 @@
     7.4    Library/AssocList.thy Library/BigO.thy Library/Binomial.thy		\
     7.5    Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy	\
     7.6    Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy	\
     7.7 -  Library/Code_Char_ord.thy Library/Code_Integer.thy Library/Code_Natural.thy			\
     7.8 -  Library/Code_Prolog.thy Tools/Predicate_Compile/code_prolog.ML	\
     7.9 -  Library/ContNotDenum.thy Library/Continuity.thy Library/Convex.thy	\
    7.10 -  Library/Countable.thy Library/Diagonalize.thy Library/Dlist.thy	\
    7.11 -  Library/Dlist_Cset.thy 						\
    7.12 +  Library/Code_Char_ord.thy Library/Code_Integer.thy			\
    7.13 +  Library/Code_Natural.thy Library/Code_Prolog.thy			\
    7.14 +  Tools/Predicate_Compile/code_prolog.ML Library/ContNotDenum.thy	\
    7.15 +  Library/Continuity.thy Library/Convex.thy Library/Countable.thy	\
    7.16 +  Library/Diagonalize.thy Library/Dlist.thy Library/Dlist_Cset.thy 	\
    7.17    Library/Efficient_Nat.thy Library/Eval_Witness.thy 			\
    7.18 -  Library/Executable_Set.thy Library/Extended_Reals.thy			\
    7.19 -  Library/Float.thy Library/Formal_Power_Series.thy			\
    7.20 -  Library/Fraction_Field.thy Library/FrechetDeriv.thy Library/Cset.thy	\
    7.21 -  Library/FuncSet.thy Library/Function_Algebras.thy			\
    7.22 -  Library/Fundamental_Theorem_Algebra.thy Library/Glbs.thy		\
    7.23 -  Library/Indicator_Function.thy Library/Infinite_Set.thy		\
    7.24 -  Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
    7.25 -  Library/LaTeXsugar.thy Library/Lattice_Algebras.thy			\
    7.26 -  Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy 	\
    7.27 -  Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy	\
    7.28 -  Library/Monad_Syntax.thy Library/More_List.thy Library/More_Set.thy	\
    7.29 -  Library/Multiset.thy Library/Nat_Bijection.thy			\
    7.30 -  Library/Nat_Infinity.thy Library/Nested_Environment.thy		\
    7.31 +  Library/Executable_Set.thy Library/Extended_Real.thy			\
    7.32 +  Library/Extended_Nat.thy Library/Float.thy				\
    7.33 +  Library/Formal_Power_Series.thy Library/Fraction_Field.thy		\
    7.34 +  Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy		\
    7.35 +  Library/Function_Algebras.thy Library/Fundamental_Theorem_Algebra.thy	\
    7.36 +  Library/Glbs.thy Library/Indicator_Function.thy			\
    7.37 +  Library/Infinite_Set.thy Library/Inner_Product.thy			\
    7.38 +  Library/Kleene_Algebra.thy Library/LaTeXsugar.thy			\
    7.39 +  Library/Lattice_Algebras.thy Library/Lattice_Syntax.thy		\
    7.40 +  Library/Library.thy Library/List_Cset.thy Library/List_Prefix.thy	\
    7.41 +  Library/List_lexord.thy Library/Mapping.thy Library/Monad_Syntax.thy	\
    7.42 +  Library/More_List.thy Library/More_Set.thy Library/Multiset.thy	\
    7.43 +  Library/Nat_Bijection.thy Library/Nested_Environment.thy		\
    7.44    Library/Numeral_Type.thy Library/OptionalSugar.thy			\
    7.45    Library/Order_Relation.thy Library/Permutation.thy			\
    7.46    Library/Permutations.thy Library/Poly_Deriv.thy			\
    7.47 @@ -481,7 +481,7 @@
    7.48    Library/Sum_of_Squares/sos_wrapper.ML					\
    7.49    Library/Sum_of_Squares/sum_of_squares.ML				\
    7.50    Library/Transitive_Closure_Table.thy Library/Univ_Poly.thy		\
    7.51 -  Library/While_Combinator.thy Library/Zorn.thy	\
    7.52 +  Library/While_Combinator.thy Library/Zorn.thy				\
    7.53    $(SRC)/Tools/adhoc_overloading.ML Library/positivstellensatz.ML	\
    7.54    Library/reflection.ML Library/reify_data.ML				\
    7.55    Library/document/root.bib Library/document/root.tex
    7.56 @@ -1201,7 +1201,7 @@
    7.57    Multivariate_Analysis/Topology_Euclidean_Space.thy			\
    7.58    Multivariate_Analysis/document/root.tex				\
    7.59    Multivariate_Analysis/normarith.ML Library/Glbs.thy			\
    7.60 -  Library/Extended_Reals.thy Library/Indicator_Function.thy		\
    7.61 +  Library/Extended_Real.thy Library/Indicator_Function.thy		\
    7.62    Library/Inner_Product.thy Library/Numeral_Type.thy Library/Convex.thy	\
    7.63    Library/FrechetDeriv.thy Library/Product_Vector.thy			\
    7.64    Library/Product_plus.thy
    7.65 @@ -1574,7 +1574,7 @@
    7.66  HOLCF-Library: HOLCF $(LOG)/HOLCF-Library.gz
    7.67  
    7.68  $(LOG)/HOLCF-Library.gz: $(OUT)/HOLCF \
    7.69 -  Library/Nat_Infinity.thy \
    7.70 +  Library/Extended_Nat.thy \
    7.71    HOLCF/Library/Defl_Bifinite.thy \
    7.72    HOLCF/Library/Bool_Discrete.thy \
    7.73    HOLCF/Library/Char_Discrete.thy \
    7.74 @@ -1628,7 +1628,7 @@
    7.75  HOLCF-FOCUS: HOLCF $(LOG)/HOLCF-FOCUS.gz
    7.76  
    7.77  $(LOG)/HOLCF-FOCUS.gz: $(OUT)/HOLCF HOLCF/FOCUS/ROOT.ML \
    7.78 -  Library/Nat_Infinity.thy \
    7.79 +  Library/Extended_Nat.thy \
    7.80    HOLCF/Library/Stream.thy \
    7.81    HOLCF/FOCUS/Fstreams.thy \
    7.82    HOLCF/FOCUS/Fstream.thy \
     8.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     8.2 +++ b/src/HOL/Library/Extended_Nat.thy	Wed Jul 20 10:48:00 2011 +0200
     8.3 @@ -0,0 +1,573 @@
     8.4 +(*  Title:      HOL/Library/Extended_Nat.thy
     8.5 +    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
     8.6 +    Contributions: David Trachtenherz, TU Muenchen
     8.7 +*)
     8.8 +
     8.9 +header {* Extended natural numbers (i.e. with infinity) *}
    8.10 +
    8.11 +theory Extended_Nat
    8.12 +imports Main
    8.13 +begin
    8.14 +
    8.15 +class infinity =
    8.16 +  fixes infinity :: "'a"
    8.17 +
    8.18 +notation (xsymbols)
    8.19 +  infinity  ("\<infinity>")
    8.20 +
    8.21 +notation (HTML output)
    8.22 +  infinity  ("\<infinity>")
    8.23 +
    8.24 +subsection {* Type definition *}
    8.25 +
    8.26 +text {*
    8.27 +  We extend the standard natural numbers by a special value indicating
    8.28 +  infinity.
    8.29 +*}
    8.30 +
    8.31 +typedef (open) enat = "UNIV :: nat option set" ..
    8.32 + 
    8.33 +definition enat :: "nat \<Rightarrow> enat" where
    8.34 +  "enat n = Abs_enat (Some n)"
    8.35 + 
    8.36 +instantiation enat :: infinity
    8.37 +begin
    8.38 +  definition "\<infinity> = Abs_enat None"
    8.39 +  instance proof qed
    8.40 +end
    8.41 + 
    8.42 +rep_datatype enat "\<infinity> :: enat"
    8.43 +proof -
    8.44 +  fix P i assume "\<And>j. P (enat j)" "P \<infinity>"
    8.45 +  then show "P i"
    8.46 +  proof induct
    8.47 +    case (Abs_enat y) then show ?case
    8.48 +      by (cases y rule: option.exhaust)
    8.49 +         (auto simp: enat_def infinity_enat_def)
    8.50 +  qed
    8.51 +qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject)
    8.52 +
    8.53 +declare [[coercion "enat::nat\<Rightarrow>enat"]]
    8.54 +
    8.55 +lemma not_Infty_eq[iff]: "(x \<noteq> \<infinity>) = (EX i. x = enat i)"
    8.56 +by (cases x) auto
    8.57 +
    8.58 +lemma not_enat_eq [iff]: "(ALL y. x ~= enat y) = (x = \<infinity>)"
    8.59 +by (cases x) auto
    8.60 +
    8.61 +primrec the_enat :: "enat \<Rightarrow> nat"
    8.62 +where "the_enat (enat n) = n"
    8.63 +
    8.64 +subsection {* Constructors and numbers *}
    8.65 +
    8.66 +instantiation enat :: "{zero, one, number}"
    8.67 +begin
    8.68 +
    8.69 +definition
    8.70 +  "0 = enat 0"
    8.71 +
    8.72 +definition
    8.73 +  [code_unfold]: "1 = enat 1"
    8.74 +
    8.75 +definition
    8.76 +  [code_unfold, code del]: "number_of k = enat (number_of k)"
    8.77 +
    8.78 +instance ..
    8.79 +
    8.80 +end
    8.81 +
    8.82 +definition iSuc :: "enat \<Rightarrow> enat" where
    8.83 +  "iSuc i = (case i of enat n \<Rightarrow> enat (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
    8.84 +
    8.85 +lemma enat_0: "enat 0 = 0"
    8.86 +  by (simp add: zero_enat_def)
    8.87 +
    8.88 +lemma enat_1: "enat 1 = 1"
    8.89 +  by (simp add: one_enat_def)
    8.90 +
    8.91 +lemma enat_number: "enat (number_of k) = number_of k"
    8.92 +  by (simp add: number_of_enat_def)
    8.93 +
    8.94 +lemma one_iSuc: "1 = iSuc 0"
    8.95 +  by (simp add: zero_enat_def one_enat_def iSuc_def)
    8.96 +
    8.97 +lemma Infty_ne_i0 [simp]: "(\<infinity>::enat) \<noteq> 0"
    8.98 +  by (simp add: zero_enat_def)
    8.99 +
   8.100 +lemma i0_ne_Infty [simp]: "0 \<noteq> (\<infinity>::enat)"
   8.101 +  by (simp add: zero_enat_def)
   8.102 +
   8.103 +lemma zero_enat_eq [simp]:
   8.104 +  "number_of k = (0\<Colon>enat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
   8.105 +  "(0\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
   8.106 +  unfolding zero_enat_def number_of_enat_def by simp_all
   8.107 +
   8.108 +lemma one_enat_eq [simp]:
   8.109 +  "number_of k = (1\<Colon>enat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
   8.110 +  "(1\<Colon>enat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
   8.111 +  unfolding one_enat_def number_of_enat_def by simp_all
   8.112 +
   8.113 +lemma zero_one_enat_neq [simp]:
   8.114 +  "\<not> 0 = (1\<Colon>enat)"
   8.115 +  "\<not> 1 = (0\<Colon>enat)"
   8.116 +  unfolding zero_enat_def one_enat_def by simp_all
   8.117 +
   8.118 +lemma Infty_ne_i1 [simp]: "(\<infinity>::enat) \<noteq> 1"
   8.119 +  by (simp add: one_enat_def)
   8.120 +
   8.121 +lemma i1_ne_Infty [simp]: "1 \<noteq> (\<infinity>::enat)"
   8.122 +  by (simp add: one_enat_def)
   8.123 +
   8.124 +lemma Infty_ne_number [simp]: "(\<infinity>::enat) \<noteq> number_of k"
   8.125 +  by (simp add: number_of_enat_def)
   8.126 +
   8.127 +lemma number_ne_Infty [simp]: "number_of k \<noteq> (\<infinity>::enat)"
   8.128 +  by (simp add: number_of_enat_def)
   8.129 +
   8.130 +lemma iSuc_enat: "iSuc (enat n) = enat (Suc n)"
   8.131 +  by (simp add: iSuc_def)
   8.132 +
   8.133 +lemma iSuc_number_of: "iSuc (number_of k) = enat (Suc (number_of k))"
   8.134 +  by (simp add: iSuc_enat number_of_enat_def)
   8.135 +
   8.136 +lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
   8.137 +  by (simp add: iSuc_def)
   8.138 +
   8.139 +lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
   8.140 +  by (simp add: iSuc_def zero_enat_def split: enat.splits)
   8.141 +
   8.142 +lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
   8.143 +  by (rule iSuc_ne_0 [symmetric])
   8.144 +
   8.145 +lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
   8.146 +  by (simp add: iSuc_def split: enat.splits)
   8.147 +
   8.148 +lemma number_of_enat_inject [simp]:
   8.149 +  "(number_of k \<Colon> enat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
   8.150 +  by (simp add: number_of_enat_def)
   8.151 +
   8.152 +
   8.153 +subsection {* Addition *}
   8.154 +
   8.155 +instantiation enat :: comm_monoid_add
   8.156 +begin
   8.157 +
   8.158 +definition [nitpick_simp]:
   8.159 +  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | enat m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | enat n \<Rightarrow> enat (m + n)))"
   8.160 +
   8.161 +lemma plus_enat_simps [simp, code]:
   8.162 +  fixes q :: enat
   8.163 +  shows "enat m + enat n = enat (m + n)"
   8.164 +    and "\<infinity> + q = \<infinity>"
   8.165 +    and "q + \<infinity> = \<infinity>"
   8.166 +  by (simp_all add: plus_enat_def split: enat.splits)
   8.167 +
   8.168 +instance proof
   8.169 +  fix n m q :: enat
   8.170 +  show "n + m + q = n + (m + q)"
   8.171 +    by (cases n, auto, cases m, auto, cases q, auto)
   8.172 +  show "n + m = m + n"
   8.173 +    by (cases n, auto, cases m, auto)
   8.174 +  show "0 + n = n"
   8.175 +    by (cases n) (simp_all add: zero_enat_def)
   8.176 +qed
   8.177 +
   8.178 +end
   8.179 +
   8.180 +lemma plus_enat_0 [simp]:
   8.181 +  "0 + (q\<Colon>enat) = q"
   8.182 +  "(q\<Colon>enat) + 0 = q"
   8.183 +  by (simp_all add: plus_enat_def zero_enat_def split: enat.splits)
   8.184 +
   8.185 +lemma plus_enat_number [simp]:
   8.186 +  "(number_of k \<Colon> enat) + number_of l = (if k < Int.Pls then number_of l
   8.187 +    else if l < Int.Pls then number_of k else number_of (k + l))"
   8.188 +  unfolding number_of_enat_def plus_enat_simps nat_arith(1) if_distrib [symmetric, of _ enat] ..
   8.189 +
   8.190 +lemma iSuc_number [simp]:
   8.191 +  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
   8.192 +  unfolding iSuc_number_of
   8.193 +  unfolding one_enat_def number_of_enat_def Suc_nat_number_of if_distrib [symmetric] ..
   8.194 +
   8.195 +lemma iSuc_plus_1:
   8.196 +  "iSuc n = n + 1"
   8.197 +  by (cases n) (simp_all add: iSuc_enat one_enat_def)
   8.198 +  
   8.199 +lemma plus_1_iSuc:
   8.200 +  "1 + q = iSuc q"
   8.201 +  "q + 1 = iSuc q"
   8.202 +by (simp_all add: iSuc_plus_1 add_ac)
   8.203 +
   8.204 +lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
   8.205 +by (simp_all add: iSuc_plus_1 add_ac)
   8.206 +
   8.207 +lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
   8.208 +by (simp only: add_commute[of m] iadd_Suc)
   8.209 +
   8.210 +lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \<and> n = 0)"
   8.211 +by (cases m, cases n, simp_all add: zero_enat_def)
   8.212 +
   8.213 +subsection {* Multiplication *}
   8.214 +
   8.215 +instantiation enat :: comm_semiring_1
   8.216 +begin
   8.217 +
   8.218 +definition times_enat_def [nitpick_simp]:
   8.219 +  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | enat m \<Rightarrow>
   8.220 +    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | enat n \<Rightarrow> enat (m * n)))"
   8.221 +
   8.222 +lemma times_enat_simps [simp, code]:
   8.223 +  "enat m * enat n = enat (m * n)"
   8.224 +  "\<infinity> * \<infinity> = (\<infinity>::enat)"
   8.225 +  "\<infinity> * enat n = (if n = 0 then 0 else \<infinity>)"
   8.226 +  "enat m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
   8.227 +  unfolding times_enat_def zero_enat_def
   8.228 +  by (simp_all split: enat.split)
   8.229 +
   8.230 +instance proof
   8.231 +  fix a b c :: enat
   8.232 +  show "(a * b) * c = a * (b * c)"
   8.233 +    unfolding times_enat_def zero_enat_def
   8.234 +    by (simp split: enat.split)
   8.235 +  show "a * b = b * a"
   8.236 +    unfolding times_enat_def zero_enat_def
   8.237 +    by (simp split: enat.split)
   8.238 +  show "1 * a = a"
   8.239 +    unfolding times_enat_def zero_enat_def one_enat_def
   8.240 +    by (simp split: enat.split)
   8.241 +  show "(a + b) * c = a * c + b * c"
   8.242 +    unfolding times_enat_def zero_enat_def
   8.243 +    by (simp split: enat.split add: left_distrib)
   8.244 +  show "0 * a = 0"
   8.245 +    unfolding times_enat_def zero_enat_def
   8.246 +    by (simp split: enat.split)
   8.247 +  show "a * 0 = 0"
   8.248 +    unfolding times_enat_def zero_enat_def
   8.249 +    by (simp split: enat.split)
   8.250 +  show "(0::enat) \<noteq> 1"
   8.251 +    unfolding zero_enat_def one_enat_def
   8.252 +    by simp
   8.253 +qed
   8.254 +
   8.255 +end
   8.256 +
   8.257 +lemma mult_iSuc: "iSuc m * n = n + m * n"
   8.258 +  unfolding iSuc_plus_1 by (simp add: algebra_simps)
   8.259 +
   8.260 +lemma mult_iSuc_right: "m * iSuc n = m + m * n"
   8.261 +  unfolding iSuc_plus_1 by (simp add: algebra_simps)
   8.262 +
   8.263 +lemma of_nat_eq_enat: "of_nat n = enat n"
   8.264 +  apply (induct n)
   8.265 +  apply (simp add: enat_0)
   8.266 +  apply (simp add: plus_1_iSuc iSuc_enat)
   8.267 +  done
   8.268 +
   8.269 +instance enat :: number_semiring
   8.270 +proof
   8.271 +  fix n show "number_of (int n) = (of_nat n :: enat)"
   8.272 +    unfolding number_of_enat_def number_of_int of_nat_id of_nat_eq_enat ..
   8.273 +qed
   8.274 +
   8.275 +instance enat :: semiring_char_0 proof
   8.276 +  have "inj enat" by (rule injI) simp
   8.277 +  then show "inj (\<lambda>n. of_nat n :: enat)" by (simp add: of_nat_eq_enat)
   8.278 +qed
   8.279 +
   8.280 +lemma imult_is_0[simp]: "((m::enat) * n = 0) = (m = 0 \<or> n = 0)"
   8.281 +by(auto simp add: times_enat_def zero_enat_def split: enat.split)
   8.282 +
   8.283 +lemma imult_is_Infty: "((a::enat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
   8.284 +by(auto simp add: times_enat_def zero_enat_def split: enat.split)
   8.285 +
   8.286 +
   8.287 +subsection {* Subtraction *}
   8.288 +
   8.289 +instantiation enat :: minus
   8.290 +begin
   8.291 +
   8.292 +definition diff_enat_def:
   8.293 +"a - b = (case a of (enat x) \<Rightarrow> (case b of (enat y) \<Rightarrow> enat (x - y) | \<infinity> \<Rightarrow> 0)
   8.294 +          | \<infinity> \<Rightarrow> \<infinity>)"
   8.295 +
   8.296 +instance ..
   8.297 +
   8.298 +end
   8.299 +
   8.300 +lemma idiff_enat_enat[simp,code]: "enat a - enat b = enat (a - b)"
   8.301 +by(simp add: diff_enat_def)
   8.302 +
   8.303 +lemma idiff_Infty[simp,code]: "\<infinity> - n = (\<infinity>::enat)"
   8.304 +by(simp add: diff_enat_def)
   8.305 +
   8.306 +lemma idiff_Infty_right[simp,code]: "enat a - \<infinity> = 0"
   8.307 +by(simp add: diff_enat_def)
   8.308 +
   8.309 +lemma idiff_0[simp]: "(0::enat) - n = 0"
   8.310 +by (cases n, simp_all add: zero_enat_def)
   8.311 +
   8.312 +lemmas idiff_enat_0[simp] = idiff_0[unfolded zero_enat_def]
   8.313 +
   8.314 +lemma idiff_0_right[simp]: "(n::enat) - 0 = n"
   8.315 +by (cases n) (simp_all add: zero_enat_def)
   8.316 +
   8.317 +lemmas idiff_enat_0_right[simp] = idiff_0_right[unfolded zero_enat_def]
   8.318 +
   8.319 +lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::enat) - n = 0"
   8.320 +by(auto simp: zero_enat_def)
   8.321 +
   8.322 +lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
   8.323 +by(simp add: iSuc_def split: enat.split)
   8.324 +
   8.325 +lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
   8.326 +by(simp add: one_enat_def iSuc_enat[symmetric] zero_enat_def[symmetric])
   8.327 +
   8.328 +(*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*)
   8.329 +
   8.330 +subsection {* Ordering *}
   8.331 +
   8.332 +instantiation enat :: linordered_ab_semigroup_add
   8.333 +begin
   8.334 +
   8.335 +definition [nitpick_simp]:
   8.336 +  "m \<le> n = (case n of enat n1 \<Rightarrow> (case m of enat m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
   8.337 +    | \<infinity> \<Rightarrow> True)"
   8.338 +
   8.339 +definition [nitpick_simp]:
   8.340 +  "m < n = (case m of enat m1 \<Rightarrow> (case n of enat n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
   8.341 +    | \<infinity> \<Rightarrow> False)"
   8.342 +
   8.343 +lemma enat_ord_simps [simp]:
   8.344 +  "enat m \<le> enat n \<longleftrightarrow> m \<le> n"
   8.345 +  "enat m < enat n \<longleftrightarrow> m < n"
   8.346 +  "q \<le> (\<infinity>::enat)"
   8.347 +  "q < (\<infinity>::enat) \<longleftrightarrow> q \<noteq> \<infinity>"
   8.348 +  "(\<infinity>::enat) \<le> q \<longleftrightarrow> q = \<infinity>"
   8.349 +  "(\<infinity>::enat) < q \<longleftrightarrow> False"
   8.350 +  by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits)
   8.351 +
   8.352 +lemma enat_ord_code [code]:
   8.353 +  "enat m \<le> enat n \<longleftrightarrow> m \<le> n"
   8.354 +  "enat m < enat n \<longleftrightarrow> m < n"
   8.355 +  "q \<le> (\<infinity>::enat) \<longleftrightarrow> True"
   8.356 +  "enat m < \<infinity> \<longleftrightarrow> True"
   8.357 +  "\<infinity> \<le> enat n \<longleftrightarrow> False"
   8.358 +  "(\<infinity>::enat) < q \<longleftrightarrow> False"
   8.359 +  by simp_all
   8.360 +
   8.361 +instance by default
   8.362 +  (auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits)
   8.363 +
   8.364 +end
   8.365 +
   8.366 +instance enat :: ordered_comm_semiring
   8.367 +proof
   8.368 +  fix a b c :: enat
   8.369 +  assume "a \<le> b" and "0 \<le> c"
   8.370 +  thus "c * a \<le> c * b"
   8.371 +    unfolding times_enat_def less_eq_enat_def zero_enat_def
   8.372 +    by (simp split: enat.splits)
   8.373 +qed
   8.374 +
   8.375 +lemma enat_ord_number [simp]:
   8.376 +  "(number_of m \<Colon> enat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
   8.377 +  "(number_of m \<Colon> enat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
   8.378 +  by (simp_all add: number_of_enat_def)
   8.379 +
   8.380 +lemma i0_lb [simp]: "(0\<Colon>enat) \<le> n"
   8.381 +  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
   8.382 +
   8.383 +lemma ile0_eq [simp]: "n \<le> (0\<Colon>enat) \<longleftrightarrow> n = 0"
   8.384 +by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
   8.385 +
   8.386 +lemma Infty_ileE [elim!]: "\<infinity> \<le> enat m \<Longrightarrow> R"
   8.387 +  by (simp add: zero_enat_def less_eq_enat_def split: enat.splits)
   8.388 +
   8.389 +lemma Infty_ilessE [elim!]: "\<infinity> < enat m \<Longrightarrow> R"
   8.390 +  by simp
   8.391 +
   8.392 +lemma not_iless0 [simp]: "\<not> n < (0\<Colon>enat)"
   8.393 +  by (simp add: zero_enat_def less_enat_def split: enat.splits)
   8.394 +
   8.395 +lemma i0_less [simp]: "(0\<Colon>enat) < n \<longleftrightarrow> n \<noteq> 0"
   8.396 +by (simp add: zero_enat_def less_enat_def split: enat.splits)
   8.397 +
   8.398 +lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
   8.399 +  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
   8.400 + 
   8.401 +lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
   8.402 +  by (simp add: iSuc_def less_enat_def split: enat.splits)
   8.403 +
   8.404 +lemma ile_iSuc [simp]: "n \<le> iSuc n"
   8.405 +  by (simp add: iSuc_def less_eq_enat_def split: enat.splits)
   8.406 +
   8.407 +lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
   8.408 +  by (simp add: zero_enat_def iSuc_def less_eq_enat_def split: enat.splits)
   8.409 +
   8.410 +lemma i0_iless_iSuc [simp]: "0 < iSuc n"
   8.411 +  by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.splits)
   8.412 +
   8.413 +lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
   8.414 +by (simp add: zero_enat_def iSuc_def less_enat_def split: enat.split)
   8.415 +
   8.416 +lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
   8.417 +  by (simp add: iSuc_def less_eq_enat_def less_enat_def split: enat.splits)
   8.418 +
   8.419 +lemma Suc_ile_eq: "enat (Suc m) \<le> n \<longleftrightarrow> enat m < n"
   8.420 +  by (cases n) auto
   8.421 +
   8.422 +lemma iless_Suc_eq [simp]: "enat m < iSuc n \<longleftrightarrow> enat m \<le> n"
   8.423 +  by (auto simp add: iSuc_def less_enat_def split: enat.splits)
   8.424 +
   8.425 +lemma imult_Infty: "(0::enat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
   8.426 +by (simp add: zero_enat_def less_enat_def split: enat.splits)
   8.427 +
   8.428 +lemma imult_Infty_right: "(0::enat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
   8.429 +by (simp add: zero_enat_def less_enat_def split: enat.splits)
   8.430 +
   8.431 +lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \<and> 0 < n)"
   8.432 +by (simp only: i0_less imult_is_0, simp)
   8.433 +
   8.434 +lemma mono_iSuc: "mono iSuc"
   8.435 +by(simp add: mono_def)
   8.436 +
   8.437 +
   8.438 +lemma min_enat_simps [simp]:
   8.439 +  "min (enat m) (enat n) = enat (min m n)"
   8.440 +  "min q 0 = 0"
   8.441 +  "min 0 q = 0"
   8.442 +  "min q (\<infinity>::enat) = q"
   8.443 +  "min (\<infinity>::enat) q = q"
   8.444 +  by (auto simp add: min_def)
   8.445 +
   8.446 +lemma max_enat_simps [simp]:
   8.447 +  "max (enat m) (enat n) = enat (max m n)"
   8.448 +  "max q 0 = q"
   8.449 +  "max 0 q = q"
   8.450 +  "max q \<infinity> = (\<infinity>::enat)"
   8.451 +  "max \<infinity> q = (\<infinity>::enat)"
   8.452 +  by (simp_all add: max_def)
   8.453 +
   8.454 +lemma enat_ile: "n \<le> enat m \<Longrightarrow> \<exists>k. n = enat k"
   8.455 +  by (cases n) simp_all
   8.456 +
   8.457 +lemma enat_iless: "n < enat m \<Longrightarrow> \<exists>k. n = enat k"
   8.458 +  by (cases n) simp_all
   8.459 +
   8.460 +lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. enat k < Y j"
   8.461 +apply (induct_tac k)
   8.462 + apply (simp (no_asm) only: enat_0)
   8.463 + apply (fast intro: le_less_trans [OF i0_lb])
   8.464 +apply (erule exE)
   8.465 +apply (drule spec)
   8.466 +apply (erule exE)
   8.467 +apply (drule ileI1)
   8.468 +apply (rule iSuc_enat [THEN subst])
   8.469 +apply (rule exI)
   8.470 +apply (erule (1) le_less_trans)
   8.471 +done
   8.472 +
   8.473 +instantiation enat :: "{bot, top}"
   8.474 +begin
   8.475 +
   8.476 +definition bot_enat :: enat where
   8.477 +  "bot_enat = 0"
   8.478 +
   8.479 +definition top_enat :: enat where
   8.480 +  "top_enat = \<infinity>"
   8.481 +
   8.482 +instance proof
   8.483 +qed (simp_all add: bot_enat_def top_enat_def)
   8.484 +
   8.485 +end
   8.486 +
   8.487 +lemma finite_enat_bounded:
   8.488 +  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> enat n"
   8.489 +  shows "finite A"
   8.490 +proof (rule finite_subset)
   8.491 +  show "finite (enat ` {..n})" by blast
   8.492 +
   8.493 +  have "A \<subseteq> {..enat n}" using le_fin by fastsimp
   8.494 +  also have "\<dots> \<subseteq> enat ` {..n}"
   8.495 +    by (rule subsetI) (case_tac x, auto)
   8.496 +  finally show "A \<subseteq> enat ` {..n}" .
   8.497 +qed
   8.498 +
   8.499 +
   8.500 +subsection {* Well-ordering *}
   8.501 +
   8.502 +lemma less_enatE:
   8.503 +  "[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P"
   8.504 +by (induct n) auto
   8.505 +
   8.506 +lemma less_InftyE:
   8.507 +  "[| n < \<infinity>; !!k. n = enat k ==> P |] ==> P"
   8.508 +by (induct n) auto
   8.509 +
   8.510 +lemma enat_less_induct:
   8.511 +  assumes prem: "!!n. \<forall>m::enat. m < n --> P m ==> P n" shows "P n"
   8.512 +proof -
   8.513 +  have P_enat: "!!k. P (enat k)"
   8.514 +    apply (rule nat_less_induct)
   8.515 +    apply (rule prem, clarify)
   8.516 +    apply (erule less_enatE, simp)
   8.517 +    done
   8.518 +  show ?thesis
   8.519 +  proof (induct n)
   8.520 +    fix nat
   8.521 +    show "P (enat nat)" by (rule P_enat)
   8.522 +  next
   8.523 +    show "P \<infinity>"
   8.524 +      apply (rule prem, clarify)
   8.525 +      apply (erule less_InftyE)
   8.526 +      apply (simp add: P_enat)
   8.527 +      done
   8.528 +  qed
   8.529 +qed
   8.530 +
   8.531 +instance enat :: wellorder
   8.532 +proof
   8.533 +  fix P and n
   8.534 +  assume hyp: "(\<And>n\<Colon>enat. (\<And>m\<Colon>enat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
   8.535 +  show "P n" by (blast intro: enat_less_induct hyp)
   8.536 +qed
   8.537 +
   8.538 +subsection {* Complete Lattice *}
   8.539 +
   8.540 +instantiation enat :: complete_lattice
   8.541 +begin
   8.542 +
   8.543 +definition inf_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
   8.544 +  "inf_enat \<equiv> min"
   8.545 +
   8.546 +definition sup_enat :: "enat \<Rightarrow> enat \<Rightarrow> enat" where
   8.547 +  "sup_enat \<equiv> max"
   8.548 +
   8.549 +definition Inf_enat :: "enat set \<Rightarrow> enat" where
   8.550 +  "Inf_enat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
   8.551 +
   8.552 +definition Sup_enat :: "enat set \<Rightarrow> enat" where
   8.553 +  "Sup_enat A \<equiv> if A = {} then 0
   8.554 +    else if finite A then Max A
   8.555 +                     else \<infinity>"
   8.556 +instance proof
   8.557 +  fix x :: "enat" and A :: "enat set"
   8.558 +  { assume "x \<in> A" then show "Inf A \<le> x"
   8.559 +      unfolding Inf_enat_def by (auto intro: Least_le) }
   8.560 +  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
   8.561 +      unfolding Inf_enat_def
   8.562 +      by (cases "A = {}") (auto intro: LeastI2_ex) }
   8.563 +  { assume "x \<in> A" then show "x \<le> Sup A"
   8.564 +      unfolding Sup_enat_def by (cases "finite A") auto }
   8.565 +  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
   8.566 +      unfolding Sup_enat_def using finite_enat_bounded by auto }
   8.567 +qed (simp_all add: inf_enat_def sup_enat_def)
   8.568 +end
   8.569 +
   8.570 +
   8.571 +subsection {* Traditional theorem names *}
   8.572 +
   8.573 +lemmas enat_defs = zero_enat_def one_enat_def number_of_enat_def iSuc_def
   8.574 +  plus_enat_def less_eq_enat_def less_enat_def
   8.575 +
   8.576 +end
     9.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     9.2 +++ b/src/HOL/Library/Extended_Real.thy	Wed Jul 20 10:48:00 2011 +0200
     9.3 @@ -0,0 +1,2543 @@
     9.4 +(*  Title:      HOL/Library/Extended_Real.thy
     9.5 +    Author:     Johannes Hölzl, TU München
     9.6 +    Author:     Robert Himmelmann, TU München
     9.7 +    Author:     Armin Heller, TU München
     9.8 +    Author:     Bogdan Grechuk, University of Edinburgh
     9.9 +*)
    9.10 +
    9.11 +header {* Extended real number line *}
    9.12 +
    9.13 +theory Extended_Real
    9.14 +  imports Complex_Main Extended_Nat
    9.15 +begin
    9.16 +
    9.17 +text {*
    9.18 +
    9.19 +For more lemmas about the extended real numbers go to
    9.20 +  @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"}
    9.21 +
    9.22 +*}
    9.23 +
    9.24 +lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
    9.25 +proof
    9.26 +  assume "{x..} = UNIV"
    9.27 +  show "x = bot"
    9.28 +  proof (rule ccontr)
    9.29 +    assume "x \<noteq> bot" then have "bot \<notin> {x..}" by (simp add: le_less)
    9.30 +    then show False using `{x..} = UNIV` by simp
    9.31 +  qed
    9.32 +qed auto
    9.33 +
    9.34 +lemma SUPR_pair:
    9.35 +  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
    9.36 +  by (rule antisym) (auto intro!: SUP_leI le_SUPI2)
    9.37 +
    9.38 +lemma INFI_pair:
    9.39 +  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
    9.40 +  by (rule antisym) (auto intro!: le_INFI INF_leI2)
    9.41 +
    9.42 +subsection {* Definition and basic properties *}
    9.43 +
    9.44 +datatype ereal = ereal real | PInfty | MInfty
    9.45 +
    9.46 +instantiation ereal :: uminus
    9.47 +begin
    9.48 +  fun uminus_ereal where
    9.49 +    "- (ereal r) = ereal (- r)"
    9.50 +  | "- PInfty = MInfty"
    9.51 +  | "- MInfty = PInfty"
    9.52 +  instance ..
    9.53 +end
    9.54 +
    9.55 +instantiation ereal :: infinity
    9.56 +begin
    9.57 +  definition "(\<infinity>::ereal) = PInfty"
    9.58 +  instance ..
    9.59 +end
    9.60 +
    9.61 +definition "ereal_of_enat n = (case n of enat n \<Rightarrow> ereal (real n) | \<infinity> \<Rightarrow> \<infinity>)"
    9.62 +
    9.63 +declare [[coercion "ereal :: real \<Rightarrow> ereal"]]
    9.64 +declare [[coercion "ereal_of_enat :: enat \<Rightarrow> ereal"]]
    9.65 +declare [[coercion "(\<lambda>n. ereal (of_nat n)) :: nat \<Rightarrow> ereal"]]
    9.66 +
    9.67 +lemma ereal_uminus_uminus[simp]:
    9.68 +  fixes a :: ereal shows "- (- a) = a"
    9.69 +  by (cases a) simp_all
    9.70 +
    9.71 +lemma
    9.72 +  shows PInfty_eq_infinity[simp]: "PInfty = \<infinity>"
    9.73 +    and MInfty_eq_minfinity[simp]: "MInfty = - \<infinity>"
    9.74 +    and MInfty_neq_PInfty[simp]: "\<infinity> \<noteq> - (\<infinity>::ereal)" "- \<infinity> \<noteq> (\<infinity>::ereal)"
    9.75 +    and MInfty_neq_ereal[simp]: "ereal r \<noteq> - \<infinity>" "- \<infinity> \<noteq> ereal r"
    9.76 +    and PInfty_neq_ereal[simp]: "ereal r \<noteq> \<infinity>" "\<infinity> \<noteq> ereal r"
    9.77 +    and PInfty_cases[simp]: "(case \<infinity> of ereal r \<Rightarrow> f r | PInfty \<Rightarrow> y | MInfty \<Rightarrow> z) = y"
    9.78 +    and MInfty_cases[simp]: "(case - \<infinity> of ereal r \<Rightarrow> f r | PInfty \<Rightarrow> y | MInfty \<Rightarrow> z) = z"
    9.79 +  by (simp_all add: infinity_ereal_def)
    9.80 +
    9.81 +lemma inj_ereal[simp]: "inj_on ereal A"
    9.82 +  unfolding inj_on_def by auto
    9.83 +
    9.84 +lemma ereal_cases[case_names real PInf MInf, cases type: ereal]:
    9.85 +  assumes "\<And>r. x = ereal r \<Longrightarrow> P"
    9.86 +  assumes "x = \<infinity> \<Longrightarrow> P"
    9.87 +  assumes "x = -\<infinity> \<Longrightarrow> P"
    9.88 +  shows P
    9.89 +  using assms by (cases x) auto
    9.90 +
    9.91 +lemmas ereal2_cases = ereal_cases[case_product ereal_cases]
    9.92 +lemmas ereal3_cases = ereal2_cases[case_product ereal_cases]
    9.93 +
    9.94 +lemma ereal_uminus_eq_iff[simp]:
    9.95 +  fixes a b :: ereal shows "-a = -b \<longleftrightarrow> a = b"
    9.96 +  by (cases rule: ereal2_cases[of a b]) simp_all
    9.97 +
    9.98 +function of_ereal :: "ereal \<Rightarrow> real" where
    9.99 +"of_ereal (ereal r) = r" |
   9.100 +"of_ereal \<infinity> = 0" |
   9.101 +"of_ereal (-\<infinity>) = 0"
   9.102 +  by (auto intro: ereal_cases)
   9.103 +termination proof qed (rule wf_empty)
   9.104 +
   9.105 +defs (overloaded)
   9.106 +  real_of_ereal_def [code_unfold]: "real \<equiv> of_ereal"
   9.107 +
   9.108 +lemma real_of_ereal[simp]:
   9.109 +    "real (- x :: ereal) = - (real x)"
   9.110 +    "real (ereal r) = r"
   9.111 +    "real (\<infinity>::ereal) = 0"
   9.112 +  by (cases x) (simp_all add: real_of_ereal_def)
   9.113 +
   9.114 +lemma range_ereal[simp]: "range ereal = UNIV - {\<infinity>, -\<infinity>}"
   9.115 +proof safe
   9.116 +  fix x assume "x \<notin> range ereal" "x \<noteq> \<infinity>"
   9.117 +  then show "x = -\<infinity>" by (cases x) auto
   9.118 +qed auto
   9.119 +
   9.120 +lemma ereal_range_uminus[simp]: "range uminus = (UNIV::ereal set)"
   9.121 +proof safe
   9.122 +  fix x :: ereal show "x \<in> range uminus" by (intro image_eqI[of _ _ "-x"]) auto
   9.123 +qed auto
   9.124 +
   9.125 +instantiation ereal :: number
   9.126 +begin
   9.127 +definition [simp]: "number_of x = ereal (number_of x)"
   9.128 +instance proof qed
   9.129 +end
   9.130 +
   9.131 +instantiation ereal :: abs
   9.132 +begin
   9.133 +  function abs_ereal where
   9.134 +    "\<bar>ereal r\<bar> = ereal \<bar>r\<bar>"
   9.135 +  | "\<bar>-\<infinity>\<bar> = (\<infinity>::ereal)"
   9.136 +  | "\<bar>\<infinity>\<bar> = (\<infinity>::ereal)"
   9.137 +  by (auto intro: ereal_cases)
   9.138 +  termination proof qed (rule wf_empty)
   9.139 +  instance ..
   9.140 +end
   9.141 +
   9.142 +lemma abs_eq_infinity_cases[elim!]: "\<lbrakk> \<bar>x :: ereal\<bar> = \<infinity> ; x = \<infinity> \<Longrightarrow> P ; x = -\<infinity> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
   9.143 +  by (cases x) auto
   9.144 +
   9.145 +lemma abs_neq_infinity_cases[elim!]: "\<lbrakk> \<bar>x :: ereal\<bar> \<noteq> \<infinity> ; \<And>r. x = ereal r \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
   9.146 +  by (cases x) auto
   9.147 +
   9.148 +lemma abs_ereal_uminus[simp]: "\<bar>- x\<bar> = \<bar>x::ereal\<bar>"
   9.149 +  by (cases x) auto
   9.150 +
   9.151 +subsubsection "Addition"
   9.152 +
   9.153 +instantiation ereal :: comm_monoid_add
   9.154 +begin
   9.155 +
   9.156 +definition "0 = ereal 0"
   9.157 +
   9.158 +function plus_ereal where
   9.159 +"ereal r + ereal p = ereal (r + p)" |
   9.160 +"\<infinity> + a = (\<infinity>::ereal)" |
   9.161 +"a + \<infinity> = (\<infinity>::ereal)" |
   9.162 +"ereal r + -\<infinity> = - \<infinity>" |
   9.163 +"-\<infinity> + ereal p = -(\<infinity>::ereal)" |
   9.164 +"-\<infinity> + -\<infinity> = -(\<infinity>::ereal)"
   9.165 +proof -
   9.166 +  case (goal1 P x)
   9.167 +  moreover then obtain a b where "x = (a, b)" by (cases x) auto
   9.168 +  ultimately show P
   9.169 +   by (cases rule: ereal2_cases[of a b]) auto
   9.170 +qed auto
   9.171 +termination proof qed (rule wf_empty)
   9.172 +
   9.173 +lemma Infty_neq_0[simp]:
   9.174 +  "(\<infinity>::ereal) \<noteq> 0" "0 \<noteq> (\<infinity>::ereal)"
   9.175 +  "-(\<infinity>::ereal) \<noteq> 0" "0 \<noteq> -(\<infinity>::ereal)"
   9.176 +  by (simp_all add: zero_ereal_def)
   9.177 +
   9.178 +lemma ereal_eq_0[simp]:
   9.179 +  "ereal r = 0 \<longleftrightarrow> r = 0"
   9.180 +  "0 = ereal r \<longleftrightarrow> r = 0"
   9.181 +  unfolding zero_ereal_def by simp_all
   9.182 +
   9.183 +instance
   9.184 +proof
   9.185 +  fix a :: ereal show "0 + a = a"
   9.186 +    by (cases a) (simp_all add: zero_ereal_def)
   9.187 +  fix b :: ereal show "a + b = b + a"
   9.188 +    by (cases rule: ereal2_cases[of a b]) simp_all
   9.189 +  fix c :: ereal show "a + b + c = a + (b + c)"
   9.190 +    by (cases rule: ereal3_cases[of a b c]) simp_all
   9.191 +qed
   9.192 +end
   9.193 +
   9.194 +lemma real_of_ereal_0[simp]: "real (0::ereal) = 0"
   9.195 +  unfolding real_of_ereal_def zero_ereal_def by simp
   9.196 +
   9.197 +lemma abs_ereal_zero[simp]: "\<bar>0\<bar> = (0::ereal)"
   9.198 +  unfolding zero_ereal_def abs_ereal.simps by simp
   9.199 +
   9.200 +lemma ereal_uminus_zero[simp]:
   9.201 +  "- 0 = (0::ereal)"
   9.202 +  by (simp add: zero_ereal_def)
   9.203 +
   9.204 +lemma ereal_uminus_zero_iff[simp]:
   9.205 +  fixes a :: ereal shows "-a = 0 \<longleftrightarrow> a = 0"
   9.206 +  by (cases a) simp_all
   9.207 +
   9.208 +lemma ereal_plus_eq_PInfty[simp]:
   9.209 +  fixes a b :: ereal shows "a + b = \<infinity> \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
   9.210 +  by (cases rule: ereal2_cases[of a b]) auto
   9.211 +
   9.212 +lemma ereal_plus_eq_MInfty[simp]:
   9.213 +  fixes a b :: ereal shows "a + b = -\<infinity> \<longleftrightarrow>
   9.214 +    (a = -\<infinity> \<or> b = -\<infinity>) \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>"
   9.215 +  by (cases rule: ereal2_cases[of a b]) auto
   9.216 +
   9.217 +lemma ereal_add_cancel_left:
   9.218 +  fixes a b :: ereal assumes "a \<noteq> -\<infinity>"
   9.219 +  shows "a + b = a + c \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
   9.220 +  using assms by (cases rule: ereal3_cases[of a b c]) auto
   9.221 +
   9.222 +lemma ereal_add_cancel_right:
   9.223 +  fixes a b :: ereal assumes "a \<noteq> -\<infinity>"
   9.224 +  shows "b + a = c + a \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
   9.225 +  using assms by (cases rule: ereal3_cases[of a b c]) auto
   9.226 +
   9.227 +lemma ereal_real:
   9.228 +  "ereal (real x) = (if \<bar>x\<bar> = \<infinity> then 0 else x)"
   9.229 +  by (cases x) simp_all
   9.230 +
   9.231 +lemma real_of_ereal_add:
   9.232 +  fixes a b :: ereal
   9.233 +  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)"
   9.234 +  by (cases rule: ereal2_cases[of a b]) auto
   9.235 +
   9.236 +subsubsection "Linear order on @{typ ereal}"
   9.237 +
   9.238 +instantiation ereal :: linorder
   9.239 +begin
   9.240 +
   9.241 +function less_ereal where
   9.242 +"   ereal x < ereal y     \<longleftrightarrow> x < y" |
   9.243 +"(\<infinity>::ereal) < a           \<longleftrightarrow> False" |
   9.244 +"         a < -(\<infinity>::ereal) \<longleftrightarrow> False" |
   9.245 +"ereal x    < \<infinity>           \<longleftrightarrow> True" |
   9.246 +"        -\<infinity> < ereal r     \<longleftrightarrow> True" |
   9.247 +"        -\<infinity> < (\<infinity>::ereal) \<longleftrightarrow> True"
   9.248 +proof -
   9.249 +  case (goal1 P x)
   9.250 +  moreover then obtain a b where "x = (a,b)" by (cases x) auto
   9.251 +  ultimately show P by (cases rule: ereal2_cases[of a b]) auto
   9.252 +qed simp_all
   9.253 +termination by (relation "{}") simp
   9.254 +
   9.255 +definition "x \<le> (y::ereal) \<longleftrightarrow> x < y \<or> x = y"
   9.256 +
   9.257 +lemma ereal_infty_less[simp]:
   9.258 +  fixes x :: ereal
   9.259 +  shows "x < \<infinity> \<longleftrightarrow> (x \<noteq> \<infinity>)"
   9.260 +    "-\<infinity> < x \<longleftrightarrow> (x \<noteq> -\<infinity>)"
   9.261 +  by (cases x, simp_all) (cases x, simp_all)
   9.262 +
   9.263 +lemma ereal_infty_less_eq[simp]:
   9.264 +  fixes x :: ereal
   9.265 +  shows "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
   9.266 +  "x \<le> -\<infinity> \<longleftrightarrow> x = -\<infinity>"
   9.267 +  by (auto simp add: less_eq_ereal_def)
   9.268 +
   9.269 +lemma ereal_less[simp]:
   9.270 +  "ereal r < 0 \<longleftrightarrow> (r < 0)"
   9.271 +  "0 < ereal r \<longleftrightarrow> (0 < r)"
   9.272 +  "0 < (\<infinity>::ereal)"
   9.273 +  "-(\<infinity>::ereal) < 0"
   9.274 +  by (simp_all add: zero_ereal_def)
   9.275 +
   9.276 +lemma ereal_less_eq[simp]:
   9.277 +  "x \<le> (\<infinity>::ereal)"
   9.278 +  "-(\<infinity>::ereal) \<le> x"
   9.279 +  "ereal r \<le> ereal p \<longleftrightarrow> r \<le> p"
   9.280 +  "ereal r \<le> 0 \<longleftrightarrow> r \<le> 0"
   9.281 +  "0 \<le> ereal r \<longleftrightarrow> 0 \<le> r"
   9.282 +  by (auto simp add: less_eq_ereal_def zero_ereal_def)
   9.283 +
   9.284 +lemma ereal_infty_less_eq2:
   9.285 +  "a \<le> b \<Longrightarrow> a = \<infinity> \<Longrightarrow> b = (\<infinity>::ereal)"
   9.286 +  "a \<le> b \<Longrightarrow> b = -\<infinity> \<Longrightarrow> a = -(\<infinity>::ereal)"
   9.287 +  by simp_all
   9.288 +
   9.289 +instance
   9.290 +proof
   9.291 +  fix x :: ereal show "x \<le> x"
   9.292 +    by (cases x) simp_all
   9.293 +  fix y :: ereal show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
   9.294 +    by (cases rule: ereal2_cases[of x y]) auto
   9.295 +  show "x \<le> y \<or> y \<le> x "
   9.296 +    by (cases rule: ereal2_cases[of x y]) auto
   9.297 +  { assume "x \<le> y" "y \<le> x" then show "x = y"
   9.298 +    by (cases rule: ereal2_cases[of x y]) auto }
   9.299 +  { fix z assume "x \<le> y" "y \<le> z" then show "x \<le> z"
   9.300 +    by (cases rule: ereal3_cases[of x y z]) auto }
   9.301 +qed
   9.302 +end
   9.303 +
   9.304 +instance ereal :: ordered_ab_semigroup_add
   9.305 +proof
   9.306 +  fix a b c :: ereal assume "a \<le> b" then show "c + a \<le> c + b"
   9.307 +    by (cases rule: ereal3_cases[of a b c]) auto
   9.308 +qed
   9.309 +
   9.310 +lemma real_of_ereal_positive_mono:
   9.311 +  fixes x y :: ereal shows "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
   9.312 +  by (cases rule: ereal2_cases[of x y]) auto
   9.313 +
   9.314 +lemma ereal_MInfty_lessI[intro, simp]:
   9.315 +  fixes a :: ereal shows "a \<noteq> -\<infinity> \<Longrightarrow> -\<infinity> < a"
   9.316 +  by (cases a) auto
   9.317 +
   9.318 +lemma ereal_less_PInfty[intro, simp]:
   9.319 +  fixes a :: ereal shows "a \<noteq> \<infinity> \<Longrightarrow> a < \<infinity>"
   9.320 +  by (cases a) auto
   9.321 +
   9.322 +lemma ereal_less_ereal_Ex:
   9.323 +  fixes a b :: ereal
   9.324 +  shows "x < ereal r \<longleftrightarrow> x = -\<infinity> \<or> (\<exists>p. p < r \<and> x = ereal p)"
   9.325 +  by (cases x) auto
   9.326 +
   9.327 +lemma less_PInf_Ex_of_nat: "x \<noteq> \<infinity> \<longleftrightarrow> (\<exists>n::nat. x < ereal (real n))"
   9.328 +proof (cases x)
   9.329 +  case (real r) then show ?thesis
   9.330 +    using reals_Archimedean2[of r] by simp
   9.331 +qed simp_all
   9.332 +
   9.333 +lemma ereal_add_mono:
   9.334 +  fixes a b c d :: ereal assumes "a \<le> b" "c \<le> d" shows "a + c \<le> b + d"
   9.335 +  using assms
   9.336 +  apply (cases a)
   9.337 +  apply (cases rule: ereal3_cases[of b c d], auto)
   9.338 +  apply (cases rule: ereal3_cases[of b c d], auto)
   9.339 +  done
   9.340 +
   9.341 +lemma ereal_minus_le_minus[simp]:
   9.342 +  fixes a b :: ereal shows "- a \<le> - b \<longleftrightarrow> b \<le> a"
   9.343 +  by (cases rule: ereal2_cases[of a b]) auto
   9.344 +
   9.345 +lemma ereal_minus_less_minus[simp]:
   9.346 +  fixes a b :: ereal shows "- a < - b \<longleftrightarrow> b < a"
   9.347 +  by (cases rule: ereal2_cases[of a b]) auto
   9.348 +
   9.349 +lemma ereal_le_real_iff:
   9.350 +  "x \<le> real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> ereal x \<le> y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x \<le> 0))"
   9.351 +  by (cases y) auto
   9.352 +
   9.353 +lemma real_le_ereal_iff:
   9.354 +  "real y \<le> x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y \<le> ereal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 \<le> x))"
   9.355 +  by (cases y) auto
   9.356 +
   9.357 +lemma ereal_less_real_iff:
   9.358 +  "x < real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> ereal x < y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x < 0))"
   9.359 +  by (cases y) auto
   9.360 +
   9.361 +lemma real_less_ereal_iff:
   9.362 +  "real y < x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y < ereal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 < x))"
   9.363 +  by (cases y) auto
   9.364 +
   9.365 +lemma real_of_ereal_pos:
   9.366 +  fixes x :: ereal shows "0 \<le> x \<Longrightarrow> 0 \<le> real x" by (cases x) auto
   9.367 +
   9.368 +lemmas real_of_ereal_ord_simps =
   9.369 +  ereal_le_real_iff real_le_ereal_iff ereal_less_real_iff real_less_ereal_iff
   9.370 +
   9.371 +lemma abs_ereal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: ereal\<bar> = x"
   9.372 +  by (cases x) auto
   9.373 +
   9.374 +lemma abs_ereal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: ereal\<bar> = -x"
   9.375 +  by (cases x) auto
   9.376 +
   9.377 +lemma abs_ereal_pos[simp]: "0 \<le> \<bar>x :: ereal\<bar>"
   9.378 +  by (cases x) auto
   9.379 +
   9.380 +lemma real_of_ereal_le_0[simp]: "real (x :: ereal) \<le> 0 \<longleftrightarrow> (x \<le> 0 \<or> x = \<infinity>)"
   9.381 +  by (cases x) auto
   9.382 +
   9.383 +lemma abs_real_of_ereal[simp]: "\<bar>real (x :: ereal)\<bar> = real \<bar>x\<bar>"
   9.384 +  by (cases x) auto
   9.385 +
   9.386 +lemma zero_less_real_of_ereal:
   9.387 +  fixes x :: ereal shows "0 < real x \<longleftrightarrow> (0 < x \<and> x \<noteq> \<infinity>)"
   9.388 +  by (cases x) auto
   9.389 +
   9.390 +lemma ereal_0_le_uminus_iff[simp]:
   9.391 +  fixes a :: ereal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
   9.392 +  by (cases rule: ereal2_cases[of a]) auto
   9.393 +
   9.394 +lemma ereal_uminus_le_0_iff[simp]:
   9.395 +  fixes a :: ereal shows "-a \<le> 0 \<longleftrightarrow> 0 \<le> a"
   9.396 +  by (cases rule: ereal2_cases[of a]) auto
   9.397 +
   9.398 +lemma ereal_dense2: "x < y \<Longrightarrow> \<exists>z. x < ereal z \<and> ereal z < y"
   9.399 +  using lt_ex gt_ex dense by (cases x y rule: ereal2_cases) auto
   9.400 +
   9.401 +lemma ereal_dense:
   9.402 +  fixes x y :: ereal assumes "x < y"
   9.403 +  shows "\<exists>z. x < z \<and> z < y"
   9.404 +  using ereal_dense2[OF `x < y`] by blast
   9.405 +
   9.406 +lemma ereal_add_strict_mono:
   9.407 +  fixes a b c d :: ereal
   9.408 +  assumes "a = b" "0 \<le> a" "a \<noteq> \<infinity>" "c < d"
   9.409 +  shows "a + c < b + d"
   9.410 +  using assms by (cases rule: ereal3_cases[case_product ereal_cases, of a b c d]) auto
   9.411 +
   9.412 +lemma ereal_less_add: 
   9.413 +  fixes a b c :: ereal shows "\<bar>a\<bar> \<noteq> \<infinity> \<Longrightarrow> c < b \<Longrightarrow> a + c < a + b"
   9.414 +  by (cases rule: ereal2_cases[of b c]) auto
   9.415 +
   9.416 +lemma ereal_uminus_eq_reorder: "- a = b \<longleftrightarrow> a = (-b::ereal)" by auto
   9.417 +
   9.418 +lemma ereal_uminus_less_reorder: "- a < b \<longleftrightarrow> -b < (a::ereal)"
   9.419 +  by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_less_minus)
   9.420 +
   9.421 +lemma ereal_uminus_le_reorder: "- a \<le> b \<longleftrightarrow> -b \<le> (a::ereal)"
   9.422 +  by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_le_minus)
   9.423 +
   9.424 +lemmas ereal_uminus_reorder =
   9.425 +  ereal_uminus_eq_reorder ereal_uminus_less_reorder ereal_uminus_le_reorder
   9.426 +
   9.427 +lemma ereal_bot:
   9.428 +  fixes x :: ereal assumes "\<And>B. x \<le> ereal B" shows "x = - \<infinity>"
   9.429 +proof (cases x)
   9.430 +  case (real r) with assms[of "r - 1"] show ?thesis by auto
   9.431 +next case PInf with assms[of 0] show ?thesis by auto
   9.432 +next case MInf then show ?thesis by simp
   9.433 +qed
   9.434 +
   9.435 +lemma ereal_top:
   9.436 +  fixes x :: ereal assumes "\<And>B. x \<ge> ereal B" shows "x = \<infinity>"
   9.437 +proof (cases x)
   9.438 +  case (real r) with assms[of "r + 1"] show ?thesis by auto
   9.439 +next case MInf with assms[of 0] show ?thesis by auto
   9.440 +next case PInf then show ?thesis by simp
   9.441 +qed
   9.442 +
   9.443 +lemma
   9.444 +  shows ereal_max[simp]: "ereal (max x y) = max (ereal x) (ereal y)"
   9.445 +    and ereal_min[simp]: "ereal (min x y) = min (ereal x) (ereal y)"
   9.446 +  by (simp_all add: min_def max_def)
   9.447 +
   9.448 +lemma ereal_max_0: "max 0 (ereal r) = ereal (max 0 r)"
   9.449 +  by (auto simp: zero_ereal_def)
   9.450 +
   9.451 +lemma
   9.452 +  fixes f :: "nat \<Rightarrow> ereal"
   9.453 +  shows incseq_uminus[simp]: "incseq (\<lambda>x. - f x) \<longleftrightarrow> decseq f"
   9.454 +  and decseq_uminus[simp]: "decseq (\<lambda>x. - f x) \<longleftrightarrow> incseq f"
   9.455 +  unfolding decseq_def incseq_def by auto
   9.456 +
   9.457 +lemma incseq_ereal: "incseq f \<Longrightarrow> incseq (\<lambda>x. ereal (f x))"
   9.458 +  unfolding incseq_def by auto
   9.459 +
   9.460 +lemma ereal_add_nonneg_nonneg:
   9.461 +  fixes a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
   9.462 +  using add_mono[of 0 a 0 b] by simp
   9.463 +
   9.464 +lemma image_eqD: "f ` A = B \<Longrightarrow> (\<forall>x\<in>A. f x \<in> B)"
   9.465 +  by auto
   9.466 +
   9.467 +lemma incseq_setsumI:
   9.468 +  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
   9.469 +  assumes "\<And>i. 0 \<le> f i"
   9.470 +  shows "incseq (\<lambda>i. setsum f {..< i})"
   9.471 +proof (intro incseq_SucI)
   9.472 +  fix n have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
   9.473 +    using assms by (rule add_left_mono)
   9.474 +  then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
   9.475 +    by auto
   9.476 +qed
   9.477 +
   9.478 +lemma incseq_setsumI2:
   9.479 +  fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
   9.480 +  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
   9.481 +  shows "incseq (\<lambda>i. \<Sum>n\<in>A. f n i)"
   9.482 +  using assms unfolding incseq_def by (auto intro: setsum_mono)
   9.483 +
   9.484 +subsubsection "Multiplication"
   9.485 +
   9.486 +instantiation ereal :: "{comm_monoid_mult, sgn}"
   9.487 +begin
   9.488 +
   9.489 +definition "1 = ereal 1"
   9.490 +
   9.491 +function sgn_ereal where
   9.492 +  "sgn (ereal r) = ereal (sgn r)"
   9.493 +| "sgn (\<infinity>::ereal) = 1"
   9.494 +| "sgn (-\<infinity>::ereal) = -1"
   9.495 +by (auto intro: ereal_cases)
   9.496 +termination proof qed (rule wf_empty)
   9.497 +
   9.498 +function times_ereal where
   9.499 +"ereal r * ereal p = ereal (r * p)" |
   9.500 +"ereal r * \<infinity> = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
   9.501 +"\<infinity> * ereal r = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
   9.502 +"ereal r * -\<infinity> = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
   9.503 +"-\<infinity> * ereal r = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
   9.504 +"(\<infinity>::ereal) * \<infinity> = \<infinity>" |
   9.505 +"-(\<infinity>::ereal) * \<infinity> = -\<infinity>" |
   9.506 +"(\<infinity>::ereal) * -\<infinity> = -\<infinity>" |
   9.507 +"-(\<infinity>::ereal) * -\<infinity> = \<infinity>"
   9.508 +proof -
   9.509 +  case (goal1 P x)
   9.510 +  moreover then obtain a b where "x = (a, b)" by (cases x) auto
   9.511 +  ultimately show P by (cases rule: ereal2_cases[of a b]) auto
   9.512 +qed simp_all
   9.513 +termination by (relation "{}") simp
   9.514 +
   9.515 +instance
   9.516 +proof
   9.517 +  fix a :: ereal show "1 * a = a"
   9.518 +    by (cases a) (simp_all add: one_ereal_def)
   9.519 +  fix b :: ereal show "a * b = b * a"
   9.520 +    by (cases rule: ereal2_cases[of a b]) simp_all
   9.521 +  fix c :: ereal show "a * b * c = a * (b * c)"
   9.522 +    by (cases rule: ereal3_cases[of a b c])
   9.523 +       (simp_all add: zero_ereal_def zero_less_mult_iff)
   9.524 +qed
   9.525 +end
   9.526 +
   9.527 +lemma real_of_ereal_le_1:
   9.528 +  fixes a :: ereal shows "a \<le> 1 \<Longrightarrow> real a \<le> 1"
   9.529 +  by (cases a) (auto simp: one_ereal_def)
   9.530 +
   9.531 +lemma abs_ereal_one[simp]: "\<bar>1\<bar> = (1::ereal)"
   9.532 +  unfolding one_ereal_def by simp
   9.533 +
   9.534 +lemma ereal_mult_zero[simp]:
   9.535 +  fixes a :: ereal shows "a * 0 = 0"
   9.536 +  by (cases a) (simp_all add: zero_ereal_def)
   9.537 +
   9.538 +lemma ereal_zero_mult[simp]:
   9.539 +  fixes a :: ereal shows "0 * a = 0"
   9.540 +  by (cases a) (simp_all add: zero_ereal_def)
   9.541 +
   9.542 +lemma ereal_m1_less_0[simp]:
   9.543 +  "-(1::ereal) < 0"
   9.544 +  by (simp add: zero_ereal_def one_ereal_def)
   9.545 +
   9.546 +lemma ereal_zero_m1[simp]:
   9.547 +  "1 \<noteq> (0::ereal)"
   9.548 +  by (simp add: zero_ereal_def one_ereal_def)
   9.549 +
   9.550 +lemma ereal_times_0[simp]:
   9.551 +  fixes x :: ereal shows "0 * x = 0"
   9.552 +  by (cases x) (auto simp: zero_ereal_def)
   9.553 +
   9.554 +lemma ereal_times[simp]:
   9.555 +  "1 \<noteq> (\<infinity>::ereal)" "(\<infinity>::ereal) \<noteq> 1"
   9.556 +  "1 \<noteq> -(\<infinity>::ereal)" "-(\<infinity>::ereal) \<noteq> 1"
   9.557 +  by (auto simp add: times_ereal_def one_ereal_def)
   9.558 +
   9.559 +lemma ereal_plus_1[simp]:
   9.560 +  "1 + ereal r = ereal (r + 1)" "ereal r + 1 = ereal (r + 1)"
   9.561 +  "1 + -(\<infinity>::ereal) = -\<infinity>" "-(\<infinity>::ereal) + 1 = -\<infinity>"
   9.562 +  unfolding one_ereal_def by auto
   9.563 +
   9.564 +lemma ereal_zero_times[simp]:
   9.565 +  fixes a b :: ereal shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
   9.566 +  by (cases rule: ereal2_cases[of a b]) auto
   9.567 +
   9.568 +lemma ereal_mult_eq_PInfty[simp]:
   9.569 +  shows "a * b = (\<infinity>::ereal) \<longleftrightarrow>
   9.570 +    (a = \<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = -\<infinity>)"
   9.571 +  by (cases rule: ereal2_cases[of a b]) auto
   9.572 +
   9.573 +lemma ereal_mult_eq_MInfty[simp]:
   9.574 +  shows "a * b = -(\<infinity>::ereal) \<longleftrightarrow>
   9.575 +    (a = \<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = -\<infinity>)"
   9.576 +  by (cases rule: ereal2_cases[of a b]) auto
   9.577 +
   9.578 +lemma ereal_0_less_1[simp]: "0 < (1::ereal)"
   9.579 +  by (simp_all add: zero_ereal_def one_ereal_def)
   9.580 +
   9.581 +lemma ereal_zero_one[simp]: "0 \<noteq> (1::ereal)"
   9.582 +  by (simp_all add: zero_ereal_def one_ereal_def)
   9.583 +
   9.584 +lemma ereal_mult_minus_left[simp]:
   9.585 +  fixes a b :: ereal shows "-a * b = - (a * b)"
   9.586 +  by (cases rule: ereal2_cases[of a b]) auto
   9.587 +
   9.588 +lemma ereal_mult_minus_right[simp]:
   9.589 +  fixes a b :: ereal shows "a * -b = - (a * b)"
   9.590 +  by (cases rule: ereal2_cases[of a b]) auto
   9.591 +
   9.592 +lemma ereal_mult_infty[simp]:
   9.593 +  "a * (\<infinity>::ereal) = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
   9.594 +  by (cases a) auto
   9.595 +
   9.596 +lemma ereal_infty_mult[simp]:
   9.597 +  "(\<infinity>::ereal) * a = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
   9.598 +  by (cases a) auto
   9.599 +
   9.600 +lemma ereal_mult_strict_right_mono:
   9.601 +  assumes "a < b" and "0 < c" "c < (\<infinity>::ereal)"
   9.602 +  shows "a * c < b * c"
   9.603 +  using assms
   9.604 +  by (cases rule: ereal3_cases[of a b c])
   9.605 +     (auto simp: zero_le_mult_iff ereal_less_PInfty)
   9.606 +
   9.607 +lemma ereal_mult_strict_left_mono:
   9.608 +  "\<lbrakk> a < b ; 0 < c ; c < (\<infinity>::ereal)\<rbrakk> \<Longrightarrow> c * a < c * b"
   9.609 +  using ereal_mult_strict_right_mono by (simp add: mult_commute[of c])
   9.610 +
   9.611 +lemma ereal_mult_right_mono:
   9.612 +  fixes a b c :: ereal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a*c \<le> b*c"
   9.613 +  using assms
   9.614 +  apply (cases "c = 0") apply simp
   9.615 +  by (cases rule: ereal3_cases[of a b c])
   9.616 +     (auto simp: zero_le_mult_iff ereal_less_PInfty)
   9.617 +
   9.618 +lemma ereal_mult_left_mono:
   9.619 +  fixes a b c :: ereal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> c * a \<le> c * b"
   9.620 +  using ereal_mult_right_mono by (simp add: mult_commute[of c])
   9.621 +
   9.622 +lemma zero_less_one_ereal[simp]: "0 \<le> (1::ereal)"
   9.623 +  by (simp add: one_ereal_def zero_ereal_def)
   9.624 +
   9.625 +lemma ereal_0_le_mult[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * (b :: ereal)"
   9.626 +  by (cases rule: ereal2_cases[of a b]) (auto simp: mult_nonneg_nonneg)
   9.627 +
   9.628 +lemma ereal_right_distrib:
   9.629 +  fixes r a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> r * (a + b) = r * a + r * b"
   9.630 +  by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps)
   9.631 +
   9.632 +lemma ereal_left_distrib:
   9.633 +  fixes r a b :: ereal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> (a + b) * r = a * r + b * r"
   9.634 +  by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps)
   9.635 +
   9.636 +lemma ereal_mult_le_0_iff:
   9.637 +  fixes a b :: ereal
   9.638 +  shows "a * b \<le> 0 \<longleftrightarrow> (0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b)"
   9.639 +  by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_le_0_iff)
   9.640 +
   9.641 +lemma ereal_zero_le_0_iff:
   9.642 +  fixes a b :: ereal
   9.643 +  shows "0 \<le> a * b \<longleftrightarrow> (0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0)"
   9.644 +  by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_le_mult_iff)
   9.645 +
   9.646 +lemma ereal_mult_less_0_iff:
   9.647 +  fixes a b :: ereal
   9.648 +  shows "a * b < 0 \<longleftrightarrow> (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b)"
   9.649 +  by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_less_0_iff)
   9.650 +
   9.651 +lemma ereal_zero_less_0_iff:
   9.652 +  fixes a b :: ereal
   9.653 +  shows "0 < a * b \<longleftrightarrow> (0 < a \<and> 0 < b) \<or> (a < 0 \<and> b < 0)"
   9.654 +  by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_less_mult_iff)
   9.655 +
   9.656 +lemma ereal_distrib:
   9.657 +  fixes a b c :: ereal
   9.658 +  assumes "a \<noteq> \<infinity> \<or> b \<noteq> -\<infinity>" "a \<noteq> -\<infinity> \<or> b \<noteq> \<infinity>" "\<bar>c\<bar> \<noteq> \<infinity>"
   9.659 +  shows "(a + b) * c = a * c + b * c"
   9.660 +  using assms
   9.661 +  by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps)
   9.662 +
   9.663 +lemma ereal_le_epsilon:
   9.664 +  fixes x y :: ereal
   9.665 +  assumes "ALL e. 0 < e --> x <= y + e"
   9.666 +  shows "x <= y"
   9.667 +proof-
   9.668 +{ assume a: "EX r. y = ereal r"
   9.669 +  from this obtain r where r_def: "y = ereal r" by auto
   9.670 +  { assume "x=(-\<infinity>)" hence ?thesis by auto }
   9.671 +  moreover
   9.672 +  { assume "~(x=(-\<infinity>))"
   9.673 +    from this obtain p where p_def: "x = ereal p"
   9.674 +    using a assms[rule_format, of 1] by (cases x) auto
   9.675 +    { fix e have "0 < e --> p <= r + e"
   9.676 +      using assms[rule_format, of "ereal e"] p_def r_def by auto }
   9.677 +    hence "p <= r" apply (subst field_le_epsilon) by auto
   9.678 +    hence ?thesis using r_def p_def by auto
   9.679 +  } ultimately have ?thesis by blast
   9.680 +}
   9.681 +moreover
   9.682 +{ assume "y=(-\<infinity>) | y=\<infinity>" hence ?thesis
   9.683 +    using assms[rule_format, of 1] by (cases x) auto
   9.684 +} ultimately show ?thesis by (cases y) auto
   9.685 +qed
   9.686 +
   9.687 +
   9.688 +lemma ereal_le_epsilon2:
   9.689 +  fixes x y :: ereal
   9.690 +  assumes "ALL e. 0 < e --> x <= y + ereal e"
   9.691 +  shows "x <= y"
   9.692 +proof-
   9.693 +{ fix e :: ereal assume "e>0"
   9.694 +  { assume "e=\<infinity>" hence "x<=y+e" by auto }
   9.695 +  moreover
   9.696 +  { assume "e~=\<infinity>"
   9.697 +    from this obtain r where "e = ereal r" using `e>0` apply (cases e) by auto
   9.698 +    hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto
   9.699 +  } ultimately have "x<=y+e" by blast
   9.700 +} from this show ?thesis using ereal_le_epsilon by auto
   9.701 +qed
   9.702 +
   9.703 +lemma ereal_le_real:
   9.704 +  fixes x y :: ereal
   9.705 +  assumes "ALL z. x <= ereal z --> y <= ereal z"
   9.706 +  shows "y <= x"
   9.707 +by (metis assms ereal_bot ereal_cases ereal_infty_less_eq ereal_less_eq linorder_le_cases)
   9.708 +
   9.709 +lemma ereal_le_ereal:
   9.710 +  fixes x y :: ereal
   9.711 +  assumes "\<And>B. B < x \<Longrightarrow> B <= y"
   9.712 +  shows "x <= y"
   9.713 +by (metis assms ereal_dense leD linorder_le_less_linear)
   9.714 +
   9.715 +lemma ereal_ge_ereal:
   9.716 +  fixes x y :: ereal
   9.717 +  assumes "ALL B. B>x --> B >= y"
   9.718 +  shows "x >= y"
   9.719 +by (metis assms ereal_dense leD linorder_le_less_linear)
   9.720 +
   9.721 +lemma setprod_ereal_0:
   9.722 +  fixes f :: "'a \<Rightarrow> ereal"
   9.723 +  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
   9.724 +proof cases
   9.725 +  assume "finite A"
   9.726 +  then show ?thesis by (induct A) auto
   9.727 +qed auto
   9.728 +
   9.729 +lemma setprod_ereal_pos:
   9.730 +  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)"
   9.731 +proof cases
   9.732 +  assume "finite I" from this pos show ?thesis by induct auto
   9.733 +qed simp
   9.734 +
   9.735 +lemma setprod_PInf:
   9.736 +  fixes f :: "'a \<Rightarrow> ereal"
   9.737 +  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
   9.738 +  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)"
   9.739 +proof cases
   9.740 +  assume "finite I" from this assms show ?thesis
   9.741 +  proof (induct I)
   9.742 +    case (insert i I)
   9.743 +    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_ereal_pos)
   9.744 +    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
   9.745 +    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
   9.746 +      using setprod_ereal_pos[of I f] pos
   9.747 +      by (cases rule: ereal2_cases[of "f i" "setprod f I"]) auto
   9.748 +    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)"
   9.749 +      using insert by (auto simp: setprod_ereal_0)
   9.750 +    finally show ?case .
   9.751 +  qed simp
   9.752 +qed simp
   9.753 +
   9.754 +lemma setprod_ereal: "(\<Prod>i\<in>A. ereal (f i)) = ereal (setprod f A)"
   9.755 +proof cases
   9.756 +  assume "finite A" then show ?thesis
   9.757 +    by induct (auto simp: one_ereal_def)
   9.758 +qed (simp add: one_ereal_def)
   9.759 +
   9.760 +subsubsection {* Power *}
   9.761 +
   9.762 +instantiation ereal :: power
   9.763 +begin
   9.764 +primrec power_ereal where
   9.765 +  "power_ereal x 0 = 1" |
   9.766 +  "power_ereal x (Suc n) = x * x ^ n"
   9.767 +instance ..
   9.768 +end
   9.769 +
   9.770 +lemma ereal_power[simp]: "(ereal x) ^ n = ereal (x^n)"
   9.771 +  by (induct n) (auto simp: one_ereal_def)
   9.772 +
   9.773 +lemma ereal_power_PInf[simp]: "(\<infinity>::ereal) ^ n = (if n = 0 then 1 else \<infinity>)"
   9.774 +  by (induct n) (auto simp: one_ereal_def)
   9.775 +
   9.776 +lemma ereal_power_uminus[simp]:
   9.777 +  fixes x :: ereal
   9.778 +  shows "(- x) ^ n = (if even n then x ^ n else - (x^n))"
   9.779 +  by (induct n) (auto simp: one_ereal_def)
   9.780 +
   9.781 +lemma ereal_power_number_of[simp]:
   9.782 +  "(number_of num :: ereal) ^ n = ereal (number_of num ^ n)"
   9.783 +  by (induct n) (auto simp: one_ereal_def)
   9.784 +
   9.785 +lemma zero_le_power_ereal[simp]:
   9.786 +  fixes a :: ereal assumes "0 \<le> a"
   9.787 +  shows "0 \<le> a ^ n"
   9.788 +  using assms by (induct n) (auto simp: ereal_zero_le_0_iff)
   9.789 +
   9.790 +subsubsection {* Subtraction *}
   9.791 +
   9.792 +lemma ereal_minus_minus_image[simp]:
   9.793 +  fixes S :: "ereal set"
   9.794 +  shows "uminus ` uminus ` S = S"
   9.795 +  by (auto simp: image_iff)
   9.796 +
   9.797 +lemma ereal_uminus_lessThan[simp]:
   9.798 +  fixes a :: ereal shows "uminus ` {..<a} = {-a<..}"
   9.799 +proof (safe intro!: image_eqI)
   9.800 +  fix x assume "-a < x"
   9.801 +  then have "- x < - (- a)" by (simp del: ereal_uminus_uminus)
   9.802 +  then show "- x < a" by simp
   9.803 +qed auto
   9.804 +
   9.805 +lemma ereal_uminus_greaterThan[simp]:
   9.806 +  "uminus ` {(a::ereal)<..} = {..<-a}"
   9.807 +  by (metis ereal_uminus_lessThan ereal_uminus_uminus
   9.808 +            ereal_minus_minus_image)
   9.809 +
   9.810 +instantiation ereal :: minus
   9.811 +begin
   9.812 +definition "x - y = x + -(y::ereal)"
   9.813 +instance ..
   9.814 +end
   9.815 +
   9.816 +lemma ereal_minus[simp]:
   9.817 +  "ereal r - ereal p = ereal (r - p)"
   9.818 +  "-\<infinity> - ereal r = -\<infinity>"
   9.819 +  "ereal r - \<infinity> = -\<infinity>"
   9.820 +  "(\<infinity>::ereal) - x = \<infinity>"
   9.821 +  "-(\<infinity>::ereal) - \<infinity> = -\<infinity>"
   9.822 +  "x - -y = x + y"
   9.823 +  "x - 0 = x"
   9.824 +  "0 - x = -x"
   9.825 +  by (simp_all add: minus_ereal_def)
   9.826 +
   9.827 +lemma ereal_x_minus_x[simp]:
   9.828 +  "x - x = (if \<bar>x\<bar> = \<infinity> then \<infinity> else 0::ereal)"
   9.829 +  by (cases x) simp_all
   9.830 +
   9.831 +lemma ereal_eq_minus_iff:
   9.832 +  fixes x y z :: ereal
   9.833 +  shows "x = z - y \<longleftrightarrow>
   9.834 +    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y = z) \<and>
   9.835 +    (y = -\<infinity> \<longrightarrow> x = \<infinity>) \<and>
   9.836 +    (y = \<infinity> \<longrightarrow> z = \<infinity> \<longrightarrow> x = \<infinity>) \<and>
   9.837 +    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>)"
   9.838 +  by (cases rule: ereal3_cases[of x y z]) auto
   9.839 +
   9.840 +lemma ereal_eq_minus:
   9.841 +  fixes x y z :: ereal
   9.842 +  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x = z - y \<longleftrightarrow> x + y = z"
   9.843 +  by (auto simp: ereal_eq_minus_iff)
   9.844 +
   9.845 +lemma ereal_less_minus_iff:
   9.846 +  fixes x y z :: ereal
   9.847 +  shows "x < z - y \<longleftrightarrow>
   9.848 +    (y = \<infinity> \<longrightarrow> z = \<infinity> \<and> x \<noteq> \<infinity>) \<and>
   9.849 +    (y = -\<infinity> \<longrightarrow> x \<noteq> \<infinity>) \<and>
   9.850 +    (\<bar>y\<bar> \<noteq> \<infinity>\<longrightarrow> x + y < z)"
   9.851 +  by (cases rule: ereal3_cases[of x y z]) auto
   9.852 +
   9.853 +lemma ereal_less_minus:
   9.854 +  fixes x y z :: ereal
   9.855 +  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x < z - y \<longleftrightarrow> x + y < z"
   9.856 +  by (auto simp: ereal_less_minus_iff)
   9.857 +
   9.858 +lemma ereal_le_minus_iff:
   9.859 +  fixes x y z :: ereal
   9.860 +  shows "x \<le> z - y \<longleftrightarrow>
   9.861 +    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>) \<and>
   9.862 +    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y \<le> z)"
   9.863 +  by (cases rule: ereal3_cases[of x y z]) auto
   9.864 +
   9.865 +lemma ereal_le_minus:
   9.866 +  fixes x y z :: ereal
   9.867 +  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x \<le> z - y \<longleftrightarrow> x + y \<le> z"
   9.868 +  by (auto simp: ereal_le_minus_iff)
   9.869 +
   9.870 +lemma ereal_minus_less_iff:
   9.871 +  fixes x y z :: ereal
   9.872 +  shows "x - y < z \<longleftrightarrow>
   9.873 +    y \<noteq> -\<infinity> \<and> (y = \<infinity> \<longrightarrow> x \<noteq> \<infinity> \<and> z \<noteq> -\<infinity>) \<and>
   9.874 +    (y \<noteq> \<infinity> \<longrightarrow> x < z + y)"
   9.875 +  by (cases rule: ereal3_cases[of x y z]) auto
   9.876 +
   9.877 +lemma ereal_minus_less:
   9.878 +  fixes x y z :: ereal
   9.879 +  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y < z \<longleftrightarrow> x < z + y"
   9.880 +  by (auto simp: ereal_minus_less_iff)
   9.881 +
   9.882 +lemma ereal_minus_le_iff:
   9.883 +  fixes x y z :: ereal
   9.884 +  shows "x - y \<le> z \<longleftrightarrow>
   9.885 +    (y = -\<infinity> \<longrightarrow> z = \<infinity>) \<and>
   9.886 +    (y = \<infinity> \<longrightarrow> x = \<infinity> \<longrightarrow> z = \<infinity>) \<and>
   9.887 +    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x \<le> z + y)"
   9.888 +  by (cases rule: ereal3_cases[of x y z]) auto
   9.889 +
   9.890 +lemma ereal_minus_le:
   9.891 +  fixes x y z :: ereal
   9.892 +  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y \<le> z \<longleftrightarrow> x \<le> z + y"
   9.893 +  by (auto simp: ereal_minus_le_iff)
   9.894 +
   9.895 +lemma ereal_minus_eq_minus_iff:
   9.896 +  fixes a b c :: ereal
   9.897 +  shows "a - b = a - c \<longleftrightarrow>
   9.898 +    b = c \<or> a = \<infinity> \<or> (a = -\<infinity> \<and> b \<noteq> -\<infinity> \<and> c \<noteq> -\<infinity>)"
   9.899 +  by (cases rule: ereal3_cases[of a b c]) auto
   9.900 +
   9.901 +lemma ereal_add_le_add_iff:
   9.902 +  fixes a b c :: ereal
   9.903 +  shows "c + a \<le> c + b \<longleftrightarrow>
   9.904 +    a \<le> b \<or> c = \<infinity> \<or> (c = -\<infinity> \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>)"
   9.905 +  by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps)
   9.906 +
   9.907 +lemma ereal_mult_le_mult_iff:
   9.908 +  fixes a b c :: ereal
   9.909 +  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)"
   9.910 +  by (cases rule: ereal3_cases[of a b c]) (simp_all add: mult_le_cancel_left)
   9.911 +
   9.912 +lemma ereal_minus_mono:
   9.913 +  fixes A B C D :: ereal assumes "A \<le> B" "D \<le> C"
   9.914 +  shows "A - C \<le> B - D"
   9.915 +  using assms
   9.916 +  by (cases rule: ereal3_cases[case_product ereal_cases, of A B C D]) simp_all
   9.917 +
   9.918 +lemma real_of_ereal_minus:
   9.919 +  fixes a b :: ereal
   9.920 +  shows "real (a - b) = (if \<bar>a\<bar> = \<infinity> \<or> \<bar>b\<bar> = \<infinity> then 0 else real a - real b)"
   9.921 +  by (cases rule: ereal2_cases[of a b]) auto
   9.922 +
   9.923 +lemma ereal_diff_positive:
   9.924 +  fixes a b :: ereal shows "a \<le> b \<Longrightarrow> 0 \<le> b - a"
   9.925 +  by (cases rule: ereal2_cases[of a b]) auto
   9.926 +
   9.927 +lemma ereal_between:
   9.928 +  fixes x e :: ereal
   9.929 +  assumes "\<bar>x\<bar> \<noteq> \<infinity>" "0 < e"
   9.930 +  shows "x - e < x" "x < x + e"
   9.931 +using assms apply (cases x, cases e) apply auto
   9.932 +using assms by (cases x, cases e) auto
   9.933 +
   9.934 +subsubsection {* Division *}
   9.935 +
   9.936 +instantiation ereal :: inverse
   9.937 +begin
   9.938 +
   9.939 +function inverse_ereal where
   9.940 +"inverse (ereal r) = (if r = 0 then \<infinity> else ereal (inverse r))" |
   9.941 +"inverse (\<infinity>::ereal) = 0" |
   9.942 +"inverse (-\<infinity>::ereal) = 0"
   9.943 +  by (auto intro: ereal_cases)
   9.944 +termination by (relation "{}") simp
   9.945 +
   9.946 +definition "x / y = x * inverse (y :: ereal)"
   9.947 +
   9.948 +instance proof qed
   9.949 +end
   9.950 +
   9.951 +lemma real_of_ereal_inverse[simp]:
   9.952 +  fixes a :: ereal
   9.953 +  shows "real (inverse a) = 1 / real a"
   9.954 +  by (cases a) (auto simp: inverse_eq_divide)
   9.955 +
   9.956 +lemma ereal_inverse[simp]:
   9.957 +  "inverse (0::ereal) = \<infinity>"
   9.958 +  "inverse (1::ereal) = 1"
   9.959 +  by (simp_all add: one_ereal_def zero_ereal_def)
   9.960 +
   9.961 +lemma ereal_divide[simp]:
   9.962 +  "ereal r / ereal p = (if p = 0 then ereal r * \<infinity> else ereal (r / p))"
   9.963 +  unfolding divide_ereal_def by (auto simp: divide_real_def)
   9.964 +
   9.965 +lemma ereal_divide_same[simp]:
   9.966 +  fixes x :: ereal shows "x / x = (if \<bar>x\<bar> = \<infinity> \<or> x = 0 then 0 else 1)"
   9.967 +  by (cases x)
   9.968 +     (simp_all add: divide_real_def divide_ereal_def one_ereal_def)
   9.969 +
   9.970 +lemma ereal_inv_inv[simp]:
   9.971 +  fixes x :: ereal shows "inverse (inverse x) = (if x \<noteq> -\<infinity> then x else \<infinity>)"
   9.972 +  by (cases x) auto
   9.973 +
   9.974 +lemma ereal_inverse_minus[simp]:
   9.975 +  fixes x :: ereal shows "inverse (- x) = (if x = 0 then \<infinity> else -inverse x)"
   9.976 +  by (cases x) simp_all
   9.977 +
   9.978 +lemma ereal_uminus_divide[simp]:
   9.979 +  fixes x y :: ereal shows "- x / y = - (x / y)"
   9.980 +  unfolding divide_ereal_def by simp
   9.981 +
   9.982 +lemma ereal_divide_Infty[simp]:
   9.983 +  fixes x :: ereal shows "x / \<infinity> = 0" "x / -\<infinity> = 0"
   9.984 +  unfolding divide_ereal_def by simp_all
   9.985 +
   9.986 +lemma ereal_divide_one[simp]:
   9.987 +  "x / 1 = (x::ereal)"
   9.988 +  unfolding divide_ereal_def by simp
   9.989 +
   9.990 +lemma ereal_divide_ereal[simp]:
   9.991 +  "\<infinity> / ereal r = (if 0 \<le> r then \<infinity> else -\<infinity>)"
   9.992 +  unfolding divide_ereal_def by simp
   9.993 +
   9.994 +lemma zero_le_divide_ereal[simp]:
   9.995 +  fixes a :: ereal assumes "0 \<le> a" "0 \<le> b"
   9.996 +  shows "0 \<le> a / b"
   9.997 +  using assms by (cases rule: ereal2_cases[of a b]) (auto simp: zero_le_divide_iff)
   9.998 +
   9.999 +lemma ereal_le_divide_pos:
  9.1000 +  fixes x y z :: ereal shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> x * y \<le> z"
  9.1001 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.1002 +
  9.1003 +lemma ereal_divide_le_pos:
  9.1004 +  fixes x y z :: ereal shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> z \<le> x * y"
  9.1005 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.1006 +
  9.1007 +lemma ereal_le_divide_neg:
  9.1008 +  fixes x y z :: ereal shows "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> z \<le> x * y"
  9.1009 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.1010 +
  9.1011 +lemma ereal_divide_le_neg:
  9.1012 +  fixes x y z :: ereal shows "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> x * y \<le> z"
  9.1013 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.1014 +
  9.1015 +lemma ereal_inverse_antimono_strict:
  9.1016 +  fixes x y :: ereal
  9.1017 +  shows "0 \<le> x \<Longrightarrow> x < y \<Longrightarrow> inverse y < inverse x"
  9.1018 +  by (cases rule: ereal2_cases[of x y]) auto
  9.1019 +
  9.1020 +lemma ereal_inverse_antimono:
  9.1021 +  fixes x y :: ereal
  9.1022 +  shows "0 \<le> x \<Longrightarrow> x <= y \<Longrightarrow> inverse y <= inverse x"
  9.1023 +  by (cases rule: ereal2_cases[of x y]) auto
  9.1024 +
  9.1025 +lemma inverse_inverse_Pinfty_iff[simp]:
  9.1026 +  fixes x :: ereal shows "inverse x = \<infinity> \<longleftrightarrow> x = 0"
  9.1027 +  by (cases x) auto
  9.1028 +
  9.1029 +lemma ereal_inverse_eq_0:
  9.1030 +  fixes x :: ereal shows "inverse x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity>"
  9.1031 +  by (cases x) auto
  9.1032 +
  9.1033 +lemma ereal_0_gt_inverse:
  9.1034 +  fixes x :: ereal shows "0 < inverse x \<longleftrightarrow> x \<noteq> \<infinity> \<and> 0 \<le> x"
  9.1035 +  by (cases x) auto
  9.1036 +
  9.1037 +lemma ereal_mult_less_right:
  9.1038 +  fixes a b c :: ereal
  9.1039 +  assumes "b * a < c * a" "0 < a" "a < \<infinity>"
  9.1040 +  shows "b < c"
  9.1041 +  using assms
  9.1042 +  by (cases rule: ereal3_cases[of a b c])
  9.1043 +     (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
  9.1044 +
  9.1045 +lemma ereal_power_divide:
  9.1046 +  fixes x y :: ereal shows "y \<noteq> 0 \<Longrightarrow> (x / y) ^ n = x^n / y^n"
  9.1047 +  by (cases rule: ereal2_cases[of x y])
  9.1048 +     (auto simp: one_ereal_def zero_ereal_def power_divide not_le
  9.1049 +                 power_less_zero_eq zero_le_power_iff)
  9.1050 +
  9.1051 +lemma ereal_le_mult_one_interval:
  9.1052 +  fixes x y :: ereal
  9.1053 +  assumes y: "y \<noteq> -\<infinity>"
  9.1054 +  assumes z: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
  9.1055 +  shows "x \<le> y"
  9.1056 +proof (cases x)
  9.1057 +  case PInf with z[of "1 / 2"] show "x \<le> y" by (simp add: one_ereal_def)
  9.1058 +next
  9.1059 +  case (real r) note r = this
  9.1060 +  show "x \<le> y"
  9.1061 +  proof (cases y)
  9.1062 +    case (real p) note p = this
  9.1063 +    have "r \<le> p"
  9.1064 +    proof (rule field_le_mult_one_interval)
  9.1065 +      fix z :: real assume "0 < z" and "z < 1"
  9.1066 +      with z[of "ereal z"]
  9.1067 +      show "z * r \<le> p" using p r by (auto simp: zero_le_mult_iff one_ereal_def)
  9.1068 +    qed
  9.1069 +    then show "x \<le> y" using p r by simp
  9.1070 +  qed (insert y, simp_all)
  9.1071 +qed simp
  9.1072 +
  9.1073 +subsection "Complete lattice"
  9.1074 +
  9.1075 +instantiation ereal :: lattice
  9.1076 +begin
  9.1077 +definition [simp]: "sup x y = (max x y :: ereal)"
  9.1078 +definition [simp]: "inf x y = (min x y :: ereal)"
  9.1079 +instance proof qed simp_all
  9.1080 +end
  9.1081 +
  9.1082 +instantiation ereal :: complete_lattice
  9.1083 +begin
  9.1084 +
  9.1085 +definition "bot = (-\<infinity>::ereal)"
  9.1086 +definition "top = (\<infinity>::ereal)"
  9.1087 +
  9.1088 +definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: ereal)"
  9.1089 +definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: ereal)"
  9.1090 +
  9.1091 +lemma ereal_complete_Sup:
  9.1092 +  fixes S :: "ereal set" assumes "S \<noteq> {}"
  9.1093 +  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
  9.1094 +proof cases
  9.1095 +  assume "\<exists>x. \<forall>a\<in>S. a \<le> ereal x"
  9.1096 +  then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> ereal y" by auto
  9.1097 +  then have "\<infinity> \<notin> S" by force
  9.1098 +  show ?thesis
  9.1099 +  proof cases
  9.1100 +    assume "S = {-\<infinity>}"
  9.1101 +    then show ?thesis by (auto intro!: exI[of _ "-\<infinity>"])
  9.1102 +  next
  9.1103 +    assume "S \<noteq> {-\<infinity>}"
  9.1104 +    with `S \<noteq> {}` `\<infinity> \<notin> S` obtain x where "x \<in> S - {-\<infinity>}" "x \<noteq> \<infinity>" by auto
  9.1105 +    with y `\<infinity> \<notin> S` have "\<forall>z\<in>real ` (S - {-\<infinity>}). z \<le> y"
  9.1106 +      by (auto simp: real_of_ereal_ord_simps)
  9.1107 +    with reals_complete2[of "real ` (S - {-\<infinity>})"] `x \<in> S - {-\<infinity>}`
  9.1108 +    obtain s where s:
  9.1109 +       "\<forall>y\<in>S - {-\<infinity>}. real y \<le> s" "\<And>z. (\<forall>y\<in>S - {-\<infinity>}. real y \<le> z) \<Longrightarrow> s \<le> z"
  9.1110 +       by auto
  9.1111 +    show ?thesis
  9.1112 +    proof (safe intro!: exI[of _ "ereal s"])
  9.1113 +      fix z assume "z \<in> S" with `\<infinity> \<notin> S` show "z \<le> ereal s"
  9.1114 +      proof (cases z)
  9.1115 +        case (real r)
  9.1116 +        then show ?thesis
  9.1117 +          using s(1)[rule_format, of z] `z \<in> S` `z = ereal r` by auto
  9.1118 +      qed auto
  9.1119 +    next
  9.1120 +      fix z assume *: "\<forall>y\<in>S. y \<le> z"
  9.1121 +      with `S \<noteq> {-\<infinity>}` `S \<noteq> {}` show "ereal s \<le> z"
  9.1122 +      proof (cases z)
  9.1123 +        case (real u)
  9.1124 +        with * have "s \<le> u"
  9.1125 +          by (intro s(2)[of u]) (auto simp: real_of_ereal_ord_simps)
  9.1126 +        then show ?thesis using real by simp
  9.1127 +      qed auto
  9.1128 +    qed
  9.1129 +  qed
  9.1130 +next
  9.1131 +  assume *: "\<not> (\<exists>x. \<forall>a\<in>S. a \<le> ereal x)"
  9.1132 +  show ?thesis
  9.1133 +  proof (safe intro!: exI[of _ \<infinity>])
  9.1134 +    fix y assume **: "\<forall>z\<in>S. z \<le> y"
  9.1135 +    with * show "\<infinity> \<le> y"
  9.1136 +    proof (cases y)
  9.1137 +      case MInf with * ** show ?thesis by (force simp: not_le)
  9.1138 +    qed auto
  9.1139 +  qed simp
  9.1140 +qed
  9.1141 +
  9.1142 +lemma ereal_complete_Inf:
  9.1143 +  fixes S :: "ereal set" assumes "S ~= {}"
  9.1144 +  shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)"
  9.1145 +proof-
  9.1146 +def S1 == "uminus ` S"
  9.1147 +hence "S1 ~= {}" using assms by auto
  9.1148 +from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)"
  9.1149 +   using ereal_complete_Sup[of S1] by auto
  9.1150 +{ fix z assume "ALL y:S. z <= y"
  9.1151 +  hence "ALL y:S1. y <= -z" unfolding S1_def by auto
  9.1152 +  hence "x <= -z" using x_def by auto
  9.1153 +  hence "z <= -x"
  9.1154 +    apply (subst ereal_uminus_uminus[symmetric])
  9.1155 +    unfolding ereal_minus_le_minus . }
  9.1156 +moreover have "(ALL y:S. -x <= y)"
  9.1157 +   using x_def unfolding S1_def
  9.1158 +   apply simp
  9.1159 +   apply (subst (3) ereal_uminus_uminus[symmetric])
  9.1160 +   unfolding ereal_minus_le_minus by simp
  9.1161 +ultimately show ?thesis by auto
  9.1162 +qed
  9.1163 +
  9.1164 +lemma ereal_complete_uminus_eq:
  9.1165 +  fixes S :: "ereal set"
  9.1166 +  shows "(\<forall>y\<in>uminus`S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>uminus`S. y \<le> z) \<longrightarrow> x \<le> z)
  9.1167 +     \<longleftrightarrow> (\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
  9.1168 +  by simp (metis ereal_minus_le_minus ereal_uminus_uminus)
  9.1169 +
  9.1170 +lemma ereal_Sup_uminus_image_eq:
  9.1171 +  fixes S :: "ereal set"
  9.1172 +  shows "Sup (uminus ` S) = - Inf S"
  9.1173 +proof cases
  9.1174 +  assume "S = {}"
  9.1175 +  moreover have "(THE x. All (op \<le> x)) = (-\<infinity>::ereal)"
  9.1176 +    by (rule the_equality) (auto intro!: ereal_bot)
  9.1177 +  moreover have "(SOME x. \<forall>y. y \<le> x) = (\<infinity>::ereal)"
  9.1178 +    by (rule some_equality) (auto intro!: ereal_top)
  9.1179 +  ultimately show ?thesis unfolding Inf_ereal_def Sup_ereal_def
  9.1180 +    Least_def Greatest_def GreatestM_def by simp
  9.1181 +next
  9.1182 +  assume "S \<noteq> {}"
  9.1183 +  with ereal_complete_Sup[of "uminus`S"]
  9.1184 +  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)"
  9.1185 +    unfolding ereal_complete_uminus_eq by auto
  9.1186 +  show "Sup (uminus ` S) = - Inf S"
  9.1187 +    unfolding Inf_ereal_def Greatest_def GreatestM_def
  9.1188 +  proof (intro someI2[of _ _ "\<lambda>x. Sup (uminus`S) = - x"])
  9.1189 +    show "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> -x)"
  9.1190 +      using x .
  9.1191 +    fix x' assume "(\<forall>y\<in>S. x' \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> x')"
  9.1192 +    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)"
  9.1193 +      unfolding ereal_complete_uminus_eq by simp
  9.1194 +    then show "Sup (uminus ` S) = -x'"
  9.1195 +      unfolding Sup_ereal_def ereal_uminus_eq_iff
  9.1196 +      by (intro Least_equality) auto
  9.1197 +  qed
  9.1198 +qed
  9.1199 +
  9.1200 +instance
  9.1201 +proof
  9.1202 +  { fix x :: ereal and A
  9.1203 +    show "bot <= x" by (cases x) (simp_all add: bot_ereal_def)
  9.1204 +    show "x <= top" by (simp add: top_ereal_def) }
  9.1205 +
  9.1206 +  { fix x :: ereal and A assume "x : A"
  9.1207 +    with ereal_complete_Sup[of A]
  9.1208 +    obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
  9.1209 +    hence "x <= s" using `x : A` by auto
  9.1210 +    also have "... = Sup A" using s unfolding Sup_ereal_def
  9.1211 +      by (auto intro!: Least_equality[symmetric])
  9.1212 +    finally show "x <= Sup A" . }
  9.1213 +  note le_Sup = this
  9.1214 +
  9.1215 +  { fix x :: ereal and A assume *: "!!z. (z : A ==> z <= x)"
  9.1216 +    show "Sup A <= x"
  9.1217 +    proof (cases "A = {}")
  9.1218 +      case True
  9.1219 +      hence "Sup A = -\<infinity>" unfolding Sup_ereal_def
  9.1220 +        by (auto intro!: Least_equality)
  9.1221 +      thus "Sup A <= x" by simp
  9.1222 +    next
  9.1223 +      case False
  9.1224 +      with ereal_complete_Sup[of A]
  9.1225 +      obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
  9.1226 +      hence "Sup A = s"
  9.1227 +        unfolding Sup_ereal_def by (auto intro!: Least_equality)
  9.1228 +      also have "s <= x" using * s by auto
  9.1229 +      finally show "Sup A <= x" .
  9.1230 +    qed }
  9.1231 +  note Sup_le = this
  9.1232 +
  9.1233 +  { fix x :: ereal and A assume "x \<in> A"
  9.1234 +    with le_Sup[of "-x" "uminus`A"] show "Inf A \<le> x"
  9.1235 +      unfolding ereal_Sup_uminus_image_eq by simp }
  9.1236 +
  9.1237 +  { fix x :: ereal and A assume *: "!!z. (z : A ==> x <= z)"
  9.1238 +    with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
  9.1239 +      unfolding ereal_Sup_uminus_image_eq by force }
  9.1240 +qed
  9.1241 +end
  9.1242 +
  9.1243 +lemma ereal_SUPR_uminus:
  9.1244 +  fixes f :: "'a => ereal"
  9.1245 +  shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
  9.1246 +  unfolding SUPR_def INFI_def
  9.1247 +  using ereal_Sup_uminus_image_eq[of "f`R"]
  9.1248 +  by (simp add: image_image)
  9.1249 +
  9.1250 +lemma ereal_INFI_uminus:
  9.1251 +  fixes f :: "'a => ereal"
  9.1252 +  shows "(INF i : R. -(f i)) = -(SUP i : R. f i)"
  9.1253 +  using ereal_SUPR_uminus[of _ "\<lambda>x. - f x"] by simp
  9.1254 +
  9.1255 +lemma ereal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::ereal set)"
  9.1256 +  using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image)
  9.1257 +
  9.1258 +lemma ereal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: ereal set)"
  9.1259 +  by (auto intro!: inj_onI)
  9.1260 +
  9.1261 +lemma ereal_image_uminus_shift:
  9.1262 +  fixes X Y :: "ereal set" shows "uminus ` X = Y \<longleftrightarrow> X = uminus ` Y"
  9.1263 +proof
  9.1264 +  assume "uminus ` X = Y"
  9.1265 +  then have "uminus ` uminus ` X = uminus ` Y"
  9.1266 +    by (simp add: inj_image_eq_iff)
  9.1267 +  then show "X = uminus ` Y" by (simp add: image_image)
  9.1268 +qed (simp add: image_image)
  9.1269 +
  9.1270 +lemma Inf_ereal_iff:
  9.1271 +  fixes z :: ereal
  9.1272 +  shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x<y) <-> Inf X < y"
  9.1273 +  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
  9.1274 +            order_less_le_trans)
  9.1275 +
  9.1276 +lemma Sup_eq_MInfty:
  9.1277 +  fixes S :: "ereal set" shows "Sup S = -\<infinity> \<longleftrightarrow> S = {} \<or> S = {-\<infinity>}"
  9.1278 +proof
  9.1279 +  assume a: "Sup S = -\<infinity>"
  9.1280 +  with complete_lattice_class.Sup_upper[of _ S]
  9.1281 +  show "S={} \<or> S={-\<infinity>}" by auto
  9.1282 +next
  9.1283 +  assume "S={} \<or> S={-\<infinity>}" then show "Sup S = -\<infinity>"
  9.1284 +    unfolding Sup_ereal_def by (auto intro!: Least_equality)
  9.1285 +qed
  9.1286 +
  9.1287 +lemma Inf_eq_PInfty:
  9.1288 +  fixes S :: "ereal set" shows "Inf S = \<infinity> \<longleftrightarrow> S = {} \<or> S = {\<infinity>}"
  9.1289 +  using Sup_eq_MInfty[of "uminus`S"]
  9.1290 +  unfolding ereal_Sup_uminus_image_eq ereal_image_uminus_shift by simp
  9.1291 +
  9.1292 +lemma Inf_eq_MInfty: 
  9.1293 +  fixes S :: "ereal set" shows "-\<infinity> \<in> S \<Longrightarrow> Inf S = -\<infinity>"
  9.1294 +  unfolding Inf_ereal_def
  9.1295 +  by (auto intro!: Greatest_equality)
  9.1296 +
  9.1297 +lemma Sup_eq_PInfty:
  9.1298 +  fixes S :: "ereal set" shows "\<infinity> \<in> S \<Longrightarrow> Sup S = \<infinity>"
  9.1299 +  unfolding Sup_ereal_def
  9.1300 +  by (auto intro!: Least_equality)
  9.1301 +
  9.1302 +lemma ereal_SUPI:
  9.1303 +  fixes x :: ereal
  9.1304 +  assumes "!!i. i : A ==> f i <= x"
  9.1305 +  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
  9.1306 +  shows "(SUP i:A. f i) = x"
  9.1307 +  unfolding SUPR_def Sup_ereal_def
  9.1308 +  using assms by (auto intro!: Least_equality)
  9.1309 +
  9.1310 +lemma ereal_INFI:
  9.1311 +  fixes x :: ereal
  9.1312 +  assumes "!!i. i : A ==> f i >= x"
  9.1313 +  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
  9.1314 +  shows "(INF i:A. f i) = x"
  9.1315 +  unfolding INFI_def Inf_ereal_def
  9.1316 +  using assms by (auto intro!: Greatest_equality)
  9.1317 +
  9.1318 +lemma Sup_ereal_close:
  9.1319 +  fixes e :: ereal
  9.1320 +  assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
  9.1321 +  shows "\<exists>x\<in>S. Sup S - e < x"
  9.1322 +  using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1])
  9.1323 +
  9.1324 +lemma Inf_ereal_close:
  9.1325 +  fixes e :: ereal assumes "\<bar>Inf X\<bar> \<noteq> \<infinity>" "0 < e"
  9.1326 +  shows "\<exists>x\<in>X. x < Inf X + e"
  9.1327 +proof (rule Inf_less_iff[THEN iffD1])
  9.1328 +  show "Inf X < Inf X + e" using assms
  9.1329 +    by (cases e) auto
  9.1330 +qed
  9.1331 +
  9.1332 +lemma Sup_eq_top_iff:
  9.1333 +  fixes A :: "'a::{complete_lattice, linorder} set"
  9.1334 +  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
  9.1335 +proof
  9.1336 +  assume *: "Sup A = top"
  9.1337 +  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
  9.1338 +  proof (intro allI impI)
  9.1339 +    fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
  9.1340 +      unfolding less_Sup_iff by auto
  9.1341 +  qed
  9.1342 +next
  9.1343 +  assume *: "\<forall>x<top. \<exists>i\<in>A. x < i"
  9.1344 +  show "Sup A = top"
  9.1345 +  proof (rule ccontr)
  9.1346 +    assume "Sup A \<noteq> top"
  9.1347 +    with top_greatest[of "Sup A"]
  9.1348 +    have "Sup A < top" unfolding le_less by auto
  9.1349 +    then have "Sup A < Sup A"
  9.1350 +      using * unfolding less_Sup_iff by auto
  9.1351 +    then show False by auto
  9.1352 +  qed
  9.1353 +qed
  9.1354 +
  9.1355 +lemma SUP_eq_top_iff:
  9.1356 +  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
  9.1357 +  shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
  9.1358 +  unfolding SUPR_def Sup_eq_top_iff by auto
  9.1359 +
  9.1360 +lemma SUP_nat_Infty: "(SUP i::nat. ereal (real i)) = \<infinity>"
  9.1361 +proof -
  9.1362 +  { fix x ::ereal assume "x \<noteq> \<infinity>"
  9.1363 +    then have "\<exists>k::nat. x < ereal (real k)"
  9.1364 +    proof (cases x)
  9.1365 +      case MInf then show ?thesis by (intro exI[of _ 0]) auto
  9.1366 +    next
  9.1367 +      case (real r)
  9.1368 +      moreover obtain k :: nat where "r < real k"
  9.1369 +        using ex_less_of_nat by (auto simp: real_eq_of_nat)
  9.1370 +      ultimately show ?thesis by auto
  9.1371 +    qed simp }
  9.1372 +  then show ?thesis
  9.1373 +    using SUP_eq_top_iff[of UNIV "\<lambda>n::nat. ereal (real n)"]
  9.1374 +    by (auto simp: top_ereal_def)
  9.1375 +qed
  9.1376 +
  9.1377 +lemma ereal_le_Sup:
  9.1378 +  fixes x :: ereal
  9.1379 +  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
  9.1380 +(is "?lhs <-> ?rhs")
  9.1381 +proof-
  9.1382 +{ assume "?rhs"
  9.1383 +  { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
  9.1384 +    from this obtain y where y_def: "(SUP i:A. f i)<y & y<x" using ereal_dense by auto
  9.1385 +    from this obtain i where "i : A & y <= f i" using `?rhs` by auto
  9.1386 +    hence "y <= (SUP i:A. f i)" using le_SUPI[of i A f] by auto
  9.1387 +    hence False using y_def by auto
  9.1388 +  } hence "?lhs" by auto
  9.1389 +}
  9.1390 +moreover
  9.1391 +{ assume "?lhs" hence "?rhs"
  9.1392 +  by (metis Collect_def Collect_mem_eq SUP_leI assms atLeastatMost_empty atLeastatMost_empty_iff
  9.1393 +      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
  9.1394 +} ultimately show ?thesis by auto
  9.1395 +qed
  9.1396 +
  9.1397 +lemma ereal_Inf_le:
  9.1398 +  fixes x :: ereal
  9.1399 +  shows "((INF i:A. f i) <= x) <-> (ALL y. x < y --> (EX i. i : A & f i <= y))"
  9.1400 +(is "?lhs <-> ?rhs")
  9.1401 +proof-
  9.1402 +{ assume "?rhs"
  9.1403 +  { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le)
  9.1404 +    from this obtain y where y_def: "x<y & y<(INF i:A. f i)" using ereal_dense by auto
  9.1405 +    from this obtain i where "i : A & f i <= y" using `?rhs` by auto
  9.1406 +    hence "(INF i:A. f i) <= y" using INF_leI[of i A f] by auto
  9.1407 +    hence False using y_def by auto
  9.1408 +  } hence "?lhs" by auto
  9.1409 +}
  9.1410 +moreover
  9.1411 +{ assume "?lhs" hence "?rhs"
  9.1412 +  by (metis Collect_def Collect_mem_eq le_INFI assms atLeastatMost_empty atLeastatMost_empty_iff
  9.1413 +      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
  9.1414 +} ultimately show ?thesis by auto
  9.1415 +qed
  9.1416 +
  9.1417 +lemma Inf_less:
  9.1418 +  fixes x :: ereal
  9.1419 +  assumes "(INF i:A. f i) < x"
  9.1420 +  shows "EX i. i : A & f i <= x"
  9.1421 +proof(rule ccontr)
  9.1422 +  assume "~ (EX i. i : A & f i <= x)"
  9.1423 +  hence "ALL i:A. f i > x" by auto
  9.1424 +  hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto
  9.1425 +  thus False using assms by auto
  9.1426 +qed
  9.1427 +
  9.1428 +lemma same_INF:
  9.1429 +  assumes "ALL e:A. f e = g e"
  9.1430 +  shows "(INF e:A. f e) = (INF e:A. g e)"
  9.1431 +proof-
  9.1432 +have "f ` A = g ` A" unfolding image_def using assms by auto
  9.1433 +thus ?thesis unfolding INFI_def by auto
  9.1434 +qed
  9.1435 +
  9.1436 +lemma same_SUP:
  9.1437 +  assumes "ALL e:A. f e = g e"
  9.1438 +  shows "(SUP e:A. f e) = (SUP e:A. g e)"
  9.1439 +proof-
  9.1440 +have "f ` A = g ` A" unfolding image_def using assms by auto
  9.1441 +thus ?thesis unfolding SUPR_def by auto
  9.1442 +qed
  9.1443 +
  9.1444 +lemma SUPR_eq:
  9.1445 +  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<le> g j"
  9.1446 +  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<le> f i"
  9.1447 +  shows "(SUP i:A. f i) = (SUP j:B. g j)"
  9.1448 +proof (intro antisym)
  9.1449 +  show "(SUP i:A. f i) \<le> (SUP j:B. g j)"
  9.1450 +    using assms by (metis SUP_leI le_SUPI2)
  9.1451 +  show "(SUP i:B. g i) \<le> (SUP j:A. f j)"
  9.1452 +    using assms by (metis SUP_leI le_SUPI2)
  9.1453 +qed
  9.1454 +
  9.1455 +lemma SUP_ereal_le_addI:
  9.1456 +  fixes f :: "'i \<Rightarrow> ereal"
  9.1457 +  assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
  9.1458 +  shows "SUPR UNIV f + y \<le> z"
  9.1459 +proof (cases y)
  9.1460 +  case (real r)
  9.1461 +  then have "\<And>i. f i \<le> z - y" using assms by (simp add: ereal_le_minus_iff)
  9.1462 +  then have "SUPR UNIV f \<le> z - y" by (rule SUP_leI)
  9.1463 +  then show ?thesis using real by (simp add: ereal_le_minus_iff)
  9.1464 +qed (insert assms, auto)
  9.1465 +
  9.1466 +lemma SUPR_ereal_add:
  9.1467 +  fixes f g :: "nat \<Rightarrow> ereal"
  9.1468 +  assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
  9.1469 +  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
  9.1470 +proof (rule ereal_SUPI)
  9.1471 +  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
  9.1472 +  have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
  9.1473 +    unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD)
  9.1474 +  { fix j
  9.1475 +    { fix i
  9.1476 +      have "f i + g j \<le> f i + g (max i j)"
  9.1477 +        using `incseq g`[THEN incseqD] by (rule add_left_mono) auto
  9.1478 +      also have "\<dots> \<le> f (max i j) + g (max i j)"
  9.1479 +        using `incseq f`[THEN incseqD] by (rule add_right_mono) auto
  9.1480 +      also have "\<dots> \<le> y" using * by auto
  9.1481 +      finally have "f i + g j \<le> y" . }
  9.1482 +    then have "SUPR UNIV f + g j \<le> y"
  9.1483 +      using assms(4)[of j] by (intro SUP_ereal_le_addI) auto
  9.1484 +    then have "g j + SUPR UNIV f \<le> y" by (simp add: ac_simps) }
  9.1485 +  then have "SUPR UNIV g + SUPR UNIV f \<le> y"
  9.1486 +    using f by (rule SUP_ereal_le_addI)
  9.1487 +  then show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
  9.1488 +qed (auto intro!: add_mono le_SUPI)
  9.1489 +
  9.1490 +lemma SUPR_ereal_add_pos:
  9.1491 +  fixes f g :: "nat \<Rightarrow> ereal"
  9.1492 +  assumes inc: "incseq f" "incseq g" and pos: "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
  9.1493 +  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
  9.1494 +proof (intro SUPR_ereal_add inc)
  9.1495 +  fix i show "f i \<noteq> -\<infinity>" "g i \<noteq> -\<infinity>" using pos[of i] by auto
  9.1496 +qed
  9.1497 +
  9.1498 +lemma SUPR_ereal_setsum:
  9.1499 +  fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> ereal"
  9.1500 +  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)" and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
  9.1501 +  shows "(SUP i. \<Sum>n\<in>A. f n i) = (\<Sum>n\<in>A. SUPR UNIV (f n))"
  9.1502 +proof cases
  9.1503 +  assume "finite A" then show ?thesis using assms
  9.1504 +    by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_ereal_add_pos)
  9.1505 +qed simp
  9.1506 +
  9.1507 +lemma SUPR_ereal_cmult:
  9.1508 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
  9.1509 +  shows "(SUP i. c * f i) = c * SUPR UNIV f"
  9.1510 +proof (rule ereal_SUPI)
  9.1511 +  fix i have "f i \<le> SUPR UNIV f" by (rule le_SUPI) auto
  9.1512 +  then show "c * f i \<le> c * SUPR UNIV f"
  9.1513 +    using `0 \<le> c` by (rule ereal_mult_left_mono)
  9.1514 +next
  9.1515 +  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c * f i \<le> y"
  9.1516 +  show "c * SUPR UNIV f \<le> y"
  9.1517 +  proof cases
  9.1518 +    assume c: "0 < c \<and> c \<noteq> \<infinity>"
  9.1519 +    with * have "SUPR UNIV f \<le> y / c"
  9.1520 +      by (intro SUP_leI) (auto simp: ereal_le_divide_pos)
  9.1521 +    with c show ?thesis
  9.1522 +      by (auto simp: ereal_le_divide_pos)
  9.1523 +  next
  9.1524 +    { assume "c = \<infinity>" have ?thesis
  9.1525 +      proof cases
  9.1526 +        assume "\<forall>i. f i = 0"
  9.1527 +        moreover then have "range f = {0}" by auto
  9.1528 +        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
  9.1529 +      next
  9.1530 +        assume "\<not> (\<forall>i. f i = 0)"
  9.1531 +        then obtain i where "f i \<noteq> 0" by auto
  9.1532 +        with *[of i] `c = \<infinity>` `0 \<le> f i` show ?thesis by (auto split: split_if_asm)
  9.1533 +      qed }
  9.1534 +    moreover assume "\<not> (0 < c \<and> c \<noteq> \<infinity>)"
  9.1535 +    ultimately show ?thesis using * `0 \<le> c` by auto
  9.1536 +  qed
  9.1537 +qed
  9.1538 +
  9.1539 +lemma SUP_PInfty:
  9.1540 +  fixes f :: "'a \<Rightarrow> ereal"
  9.1541 +  assumes "\<And>n::nat. \<exists>i\<in>A. ereal (real n) \<le> f i"
  9.1542 +  shows "(SUP i:A. f i) = \<infinity>"
  9.1543 +  unfolding SUPR_def Sup_eq_top_iff[where 'a=ereal, unfolded top_ereal_def]
  9.1544 +  apply simp
  9.1545 +proof safe
  9.1546 +  fix x :: ereal assume "x \<noteq> \<infinity>"
  9.1547 +  show "\<exists>i\<in>A. x < f i"
  9.1548 +  proof (cases x)
  9.1549 +    case PInf with `x \<noteq> \<infinity>` show ?thesis by simp
  9.1550 +  next
  9.1551 +    case MInf with assms[of "0"] show ?thesis by force
  9.1552 +  next
  9.1553 +    case (real r)
  9.1554 +    with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < ereal (real n)" by auto
  9.1555 +    moreover from assms[of n] guess i ..
  9.1556 +    ultimately show ?thesis
  9.1557 +      by (auto intro!: bexI[of _ i])
  9.1558 +  qed
  9.1559 +qed
  9.1560 +
  9.1561 +lemma Sup_countable_SUPR:
  9.1562 +  assumes "A \<noteq> {}"
  9.1563 +  shows "\<exists>f::nat \<Rightarrow> ereal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
  9.1564 +proof (cases "Sup A")
  9.1565 +  case (real r)
  9.1566 +  have "\<forall>n::nat. \<exists>x. x \<in> A \<and> Sup A < x + 1 / ereal (real n)"
  9.1567 +  proof
  9.1568 +    fix n ::nat have "\<exists>x\<in>A. Sup A - 1 / ereal (real n) < x"
  9.1569 +      using assms real by (intro Sup_ereal_close) (auto simp: one_ereal_def)
  9.1570 +    then guess x ..
  9.1571 +    then show "\<exists>x. x \<in> A \<and> Sup A < x + 1 / ereal (real n)"
  9.1572 +      by (auto intro!: exI[of _ x] simp: ereal_minus_less_iff)
  9.1573 +  qed
  9.1574 +  from choice[OF this] guess f .. note f = this
  9.1575 +  have "SUPR UNIV f = Sup A"
  9.1576 +  proof (rule ereal_SUPI)
  9.1577 +    fix i show "f i \<le> Sup A" using f
  9.1578 +      by (auto intro!: complete_lattice_class.Sup_upper)
  9.1579 +  next
  9.1580 +    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
  9.1581 +    show "Sup A \<le> y"
  9.1582 +    proof (rule ereal_le_epsilon, intro allI impI)
  9.1583 +      fix e :: ereal assume "0 < e"
  9.1584 +      show "Sup A \<le> y + e"
  9.1585 +      proof (cases e)
  9.1586 +        case (real r)
  9.1587 +        hence "0 < r" using `0 < e` by auto
  9.1588 +        then obtain n ::nat where *: "1 / real n < r" "0 < n"
  9.1589 +          using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
  9.1590 +        have "Sup A \<le> f n + 1 / ereal (real n)" using f[THEN spec, of n] by auto
  9.1591 +        also have "1 / ereal (real n) \<le> e" using real * by (auto simp: one_ereal_def )
  9.1592 +        with bound have "f n + 1 / ereal (real n) \<le> y + e" by (rule add_mono) simp
  9.1593 +        finally show "Sup A \<le> y + e" .
  9.1594 +      qed (insert `0 < e`, auto)
  9.1595 +    qed
  9.1596 +  qed
  9.1597 +  with f show ?thesis by (auto intro!: exI[of _ f])
  9.1598 +next
  9.1599 +  case PInf
  9.1600 +  from `A \<noteq> {}` obtain x where "x \<in> A" by auto
  9.1601 +  show ?thesis
  9.1602 +  proof cases
  9.1603 +    assume "\<infinity> \<in> A"
  9.1604 +    moreover then have "\<infinity> \<le> Sup A" by (intro complete_lattice_class.Sup_upper)
  9.1605 +    ultimately show ?thesis by (auto intro!: exI[of _ "\<lambda>x. \<infinity>"])
  9.1606 +  next
  9.1607 +    assume "\<infinity> \<notin> A"
  9.1608 +    have "\<exists>x\<in>A. 0 \<le> x"
  9.1609 +      by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least ereal_infty_less_eq2 linorder_linear)
  9.1610 +    then obtain x where "x \<in> A" "0 \<le> x" by auto
  9.1611 +    have "\<forall>n::nat. \<exists>f. f \<in> A \<and> x + ereal (real n) \<le> f"
  9.1612 +    proof (rule ccontr)
  9.1613 +      assume "\<not> ?thesis"
  9.1614 +      then have "\<exists>n::nat. Sup A \<le> x + ereal (real n)"
  9.1615 +        by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le)
  9.1616 +      then show False using `x \<in> A` `\<infinity> \<notin> A` PInf
  9.1617 +        by(cases x) auto
  9.1618 +    qed
  9.1619 +    from choice[OF this] guess f .. note f = this
  9.1620 +    have "SUPR UNIV f = \<infinity>"
  9.1621 +    proof (rule SUP_PInfty)
  9.1622 +      fix n :: nat show "\<exists>i\<in>UNIV. ereal (real n) \<le> f i"
  9.1623 +        using f[THEN spec, of n] `0 \<le> x`
  9.1624 +        by (cases rule: ereal2_cases[of "f n" x]) (auto intro!: exI[of _ n])
  9.1625 +    qed
  9.1626 +    then show ?thesis using f PInf by (auto intro!: exI[of _ f])
  9.1627 +  qed
  9.1628 +next
  9.1629 +  case MInf
  9.1630 +  with `A \<noteq> {}` have "A = {-\<infinity>}" by (auto simp: Sup_eq_MInfty)
  9.1631 +  then show ?thesis using MInf by (auto intro!: exI[of _ "\<lambda>x. -\<infinity>"])
  9.1632 +qed
  9.1633 +
  9.1634 +lemma SUPR_countable_SUPR:
  9.1635 +  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> ereal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
  9.1636 +  using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def)
  9.1637 +
  9.1638 +lemma Sup_ereal_cadd:
  9.1639 +  fixes A :: "ereal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
  9.1640 +  shows "Sup ((\<lambda>x. a + x) ` A) = a + Sup A"
  9.1641 +proof (rule antisym)
  9.1642 +  have *: "\<And>a::ereal. \<And>A. Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A"
  9.1643 +    by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
  9.1644 +  then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
  9.1645 +  show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
  9.1646 +  proof (cases a)
  9.1647 +    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
  9.1648 +  next
  9.1649 +    case (real r)
  9.1650 +    then have **: "op + (- a) ` op + a ` A = A"
  9.1651 +      by (auto simp: image_iff ac_simps zero_ereal_def[symmetric])
  9.1652 +    from *[of "-a" "(\<lambda>x. a + x) ` A"] real show ?thesis unfolding **
  9.1653 +      by (cases rule: ereal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto
  9.1654 +  qed (insert `a \<noteq> -\<infinity>`, auto)
  9.1655 +qed
  9.1656 +
  9.1657 +lemma Sup_ereal_cminus:
  9.1658 +  fixes A :: "ereal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
  9.1659 +  shows "Sup ((\<lambda>x. a - x) ` A) = a - Inf A"
  9.1660 +  using Sup_ereal_cadd[of "uminus ` A" a] assms
  9.1661 +  by (simp add: comp_def image_image minus_ereal_def
  9.1662 +                 ereal_Sup_uminus_image_eq)
  9.1663 +
  9.1664 +lemma SUPR_ereal_cminus:
  9.1665 +  fixes f :: "'i \<Rightarrow> ereal"
  9.1666 +  fixes A assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
  9.1667 +  shows "(SUP x:A. a - f x) = a - (INF x:A. f x)"
  9.1668 +  using Sup_ereal_cminus[of "f`A" a] assms
  9.1669 +  unfolding SUPR_def INFI_def image_image by auto
  9.1670 +
  9.1671 +lemma Inf_ereal_cminus:
  9.1672 +  fixes A :: "ereal set" assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
  9.1673 +  shows "Inf ((\<lambda>x. a - x) ` A) = a - Sup A"
  9.1674 +proof -
  9.1675 +  { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto }
  9.1676 +  moreover then have "(\<lambda>x. -a - x)`uminus`A = uminus ` (\<lambda>x. a - x) ` A"
  9.1677 +    by (auto simp: image_image)
  9.1678 +  ultimately show ?thesis
  9.1679 +    using Sup_ereal_cminus[of "uminus ` A" "-a"] assms
  9.1680 +    by (auto simp add: ereal_Sup_uminus_image_eq ereal_Inf_uminus_image_eq)
  9.1681 +qed
  9.1682 +
  9.1683 +lemma INFI_ereal_cminus:
  9.1684 +  fixes a :: ereal assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
  9.1685 +  shows "(INF x:A. a - f x) = a - (SUP x:A. f x)"
  9.1686 +  using Inf_ereal_cminus[of "f`A" a] assms
  9.1687 +  unfolding SUPR_def INFI_def image_image
  9.1688 +  by auto
  9.1689 +
  9.1690 +lemma uminus_ereal_add_uminus_uminus:
  9.1691 +  fixes a b :: ereal shows "a \<noteq> \<infinity> \<Longrightarrow> b \<noteq> \<infinity> \<Longrightarrow> - (- a + - b) = a + b"
  9.1692 +  by (cases rule: ereal2_cases[of a b]) auto
  9.1693 +
  9.1694 +lemma INFI_ereal_add:
  9.1695 +  fixes f :: "nat \<Rightarrow> ereal"
  9.1696 +  assumes "decseq f" "decseq g" and fin: "\<And>i. f i \<noteq> \<infinity>" "\<And>i. g i \<noteq> \<infinity>"
  9.1697 +  shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g"
  9.1698 +proof -
  9.1699 +  have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
  9.1700 +    using assms unfolding INF_less_iff by auto
  9.1701 +  { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i"
  9.1702 +      by (rule uminus_ereal_add_uminus_uminus) }
  9.1703 +  then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
  9.1704 +    by simp
  9.1705 +  also have "\<dots> = INFI UNIV f + INFI UNIV g"
  9.1706 +    unfolding ereal_INFI_uminus
  9.1707 +    using assms INF_less
  9.1708 +    by (subst SUPR_ereal_add)
  9.1709 +       (auto simp: ereal_SUPR_uminus intro!: uminus_ereal_add_uminus_uminus)
  9.1710 +  finally show ?thesis .
  9.1711 +qed
  9.1712 +
  9.1713 +subsection "Limits on @{typ ereal}"
  9.1714 +
  9.1715 +subsubsection "Topological space"
  9.1716 +
  9.1717 +instantiation ereal :: topological_space
  9.1718 +begin
  9.1719 +
  9.1720 +definition "open A \<longleftrightarrow> open (ereal -` A)
  9.1721 +       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {ereal x <..} \<subseteq> A))
  9.1722 +       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A))"
  9.1723 +
  9.1724 +lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {ereal x<..} \<subseteq> A)"
  9.1725 +  unfolding open_ereal_def by auto
  9.1726 +
  9.1727 +lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<ereal x} \<subseteq> A)"
  9.1728 +  unfolding open_ereal_def by auto
  9.1729 +
  9.1730 +lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{ereal x<..} \<subseteq> A"
  9.1731 +  using open_PInfty[OF assms] by auto
  9.1732 +
  9.1733 +lemma open_MInfty2: assumes "open A" "-\<infinity> \<in> A" obtains x where "{..<ereal x} \<subseteq> A"
  9.1734 +  using open_MInfty[OF assms] by auto
  9.1735 +
  9.1736 +lemma ereal_openE: assumes "open A" obtains x y where
  9.1737 +  "open (ereal -` A)"
  9.1738 +  "\<infinity> \<in> A \<Longrightarrow> {ereal x<..} \<subseteq> A"
  9.1739 +  "-\<infinity> \<in> A \<Longrightarrow> {..<ereal y} \<subseteq> A"
  9.1740 +  using assms open_ereal_def by auto
  9.1741 +
  9.1742 +instance
  9.1743 +proof
  9.1744 +  let ?U = "UNIV::ereal set"
  9.1745 +  show "open ?U" unfolding open_ereal_def
  9.1746 +    by (auto intro!: exI[of _ 0])
  9.1747 +next
  9.1748 +  fix S T::"ereal set" assume "open S" and "open T"
  9.1749 +  from `open S`[THEN ereal_openE] guess xS yS .
  9.1750 +  moreover from `open T`[THEN ereal_openE] guess xT yT .
  9.1751 +  ultimately have
  9.1752 +    "open (ereal -` (S \<inter> T))"
  9.1753 +    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {ereal (max xS xT) <..} \<subseteq> S \<inter> T"
  9.1754 +    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< ereal (min yS yT)} \<subseteq> S \<inter> T"
  9.1755 +    by auto
  9.1756 +  then show "open (S Int T)" unfolding open_ereal_def by blast
  9.1757 +next
  9.1758 +  fix K :: "ereal set set" assume "\<forall>S\<in>K. open S"
  9.1759 +  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (ereal -` S) \<and>
  9.1760 +    (\<infinity> \<in> S \<longrightarrow> {ereal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< ereal y} \<subseteq> S)"
  9.1761 +    by (auto simp: open_ereal_def)
  9.1762 +  then show "open (Union K)" unfolding open_ereal_def
  9.1763 +  proof (intro conjI impI)
  9.1764 +    show "open (ereal -` \<Union>K)"
  9.1765 +      using *[THEN choice] by (auto simp: vimage_Union)
  9.1766 +  qed ((metis UnionE Union_upper subset_trans *)+)
  9.1767 +qed
  9.1768 +end
  9.1769 +
  9.1770 +lemma open_ereal: "open S \<Longrightarrow> open (ereal ` S)"
  9.1771 +  by (auto simp: inj_vimage_image_eq open_ereal_def)
  9.1772 +
  9.1773 +lemma open_ereal_vimage: "open S \<Longrightarrow> open (ereal -` S)"
  9.1774 +  unfolding open_ereal_def by auto
  9.1775 +
  9.1776 +lemma open_ereal_lessThan[intro, simp]: "open {..< a :: ereal}"
  9.1777 +proof -
  9.1778 +  have "\<And>x. ereal -` {..<ereal x} = {..< x}"
  9.1779 +    "ereal -` {..< \<infinity>} = UNIV" "ereal -` {..< -\<infinity>} = {}" by auto
  9.1780 +  then show ?thesis by (cases a) (auto simp: open_ereal_def)
  9.1781 +qed
  9.1782 +
  9.1783 +lemma open_ereal_greaterThan[intro, simp]:
  9.1784 +  "open {a :: ereal <..}"
  9.1785 +proof -
  9.1786 +  have "\<And>x. ereal -` {ereal x<..} = {x<..}"
  9.1787 +    "ereal -` {\<infinity><..} = {}" "ereal -` {-\<infinity><..} = UNIV" by auto
  9.1788 +  then show ?thesis by (cases a) (auto simp: open_ereal_def)
  9.1789 +qed
  9.1790 +
  9.1791 +lemma ereal_open_greaterThanLessThan[intro, simp]: "open {a::ereal <..< b}"
  9.1792 +  unfolding greaterThanLessThan_def by auto
  9.1793 +
  9.1794 +lemma closed_ereal_atLeast[simp, intro]: "closed {a :: ereal ..}"
  9.1795 +proof -
  9.1796 +  have "- {a ..} = {..< a}" by auto
  9.1797 +  then show "closed {a ..}"
  9.1798 +    unfolding closed_def using open_ereal_lessThan by auto
  9.1799 +qed
  9.1800 +
  9.1801 +lemma closed_ereal_atMost[simp, intro]: "closed {.. b :: ereal}"
  9.1802 +proof -
  9.1803 +  have "- {.. b} = {b <..}" by auto
  9.1804 +  then show "closed {.. b}"
  9.1805 +    unfolding closed_def using open_ereal_greaterThan by auto
  9.1806 +qed
  9.1807 +
  9.1808 +lemma closed_ereal_atLeastAtMost[simp, intro]:
  9.1809 +  shows "closed {a :: ereal .. b}"
  9.1810 +  unfolding atLeastAtMost_def by auto
  9.1811 +
  9.1812 +lemma closed_ereal_singleton:
  9.1813 +  "closed {a :: ereal}"
  9.1814 +by (metis atLeastAtMost_singleton closed_ereal_atLeastAtMost)
  9.1815 +
  9.1816 +lemma ereal_open_cont_interval:
  9.1817 +  fixes S :: "ereal set"
  9.1818 +  assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
  9.1819 +  obtains e where "e>0" "{x-e <..< x+e} \<subseteq> S"
  9.1820 +proof-
  9.1821 +  from `open S` have "open (ereal -` S)" by (rule ereal_openE)
  9.1822 +  then obtain e where "0 < e" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> ereal y \<in> S"
  9.1823 +    using assms unfolding open_dist by force
  9.1824 +  show thesis
  9.1825 +  proof (intro that subsetI)
  9.1826 +    show "0 < ereal e" using `0 < e` by auto
  9.1827 +    fix y assume "y \<in> {x - ereal e<..<x + ereal e}"
  9.1828 +    with assms obtain t where "y = ereal t" "dist t (real x) < e"
  9.1829 +      apply (cases y) by (auto simp: dist_real_def)
  9.1830 +    then show "y \<in> S" using e[of t] by auto
  9.1831 +  qed
  9.1832 +qed
  9.1833 +
  9.1834 +lemma ereal_open_cont_interval2:
  9.1835 +  fixes S :: "ereal set"
  9.1836 +  assumes "open S" "x \<in> S" and x: "\<bar>x\<bar> \<noteq> \<infinity>"
  9.1837 +  obtains a b where "a < x" "x < b" "{a <..< b} \<subseteq> S"
  9.1838 +proof-
  9.1839 +  guess e using ereal_open_cont_interval[OF assms] .
  9.1840 +  with that[of "x-e" "x+e"] ereal_between[OF x, of e]
  9.1841 +  show thesis by auto
  9.1842 +qed
  9.1843 +
  9.1844 +instance ereal :: t2_space
  9.1845 +proof
  9.1846 +  fix x y :: ereal assume "x ~= y"
  9.1847 +  let "?P x (y::ereal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
  9.1848 +
  9.1849 +  { fix x y :: ereal assume "x < y"
  9.1850 +    from ereal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
  9.1851 +    have "?P x y"
  9.1852 +      apply (rule exI[of _ "{..<z}"])
  9.1853 +      apply (rule exI[of _ "{z<..}"])
  9.1854 +      using z by auto }
  9.1855 +  note * = this
  9.1856 +
  9.1857 +  from `x ~= y`
  9.1858 +  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
  9.1859 +  proof (cases rule: linorder_cases)
  9.1860 +    assume "x = y" with `x ~= y` show ?thesis by simp
  9.1861 +  next assume "x < y" from *[OF this] show ?thesis by auto
  9.1862 +  next assume "y < x" from *[OF this] show ?thesis by auto
  9.1863 +  qed
  9.1864 +qed
  9.1865 +
  9.1866 +subsubsection {* Convergent sequences *}
  9.1867 +
  9.1868 +lemma lim_ereal[simp]:
  9.1869 +  "((\<lambda>n. ereal (f n)) ---> ereal x) net \<longleftrightarrow> (f ---> x) net" (is "?l = ?r")
  9.1870 +proof (intro iffI topological_tendstoI)
  9.1871 +  fix S assume "?l" "open S" "x \<in> S"
  9.1872 +  then show "eventually (\<lambda>x. f x \<in> S) net"
  9.1873 +    using `?l`[THEN topological_tendstoD, OF open_ereal, OF `open S`]
  9.1874 +    by (simp add: inj_image_mem_iff)
  9.1875 +next
  9.1876 +  fix S assume "?r" "open S" "ereal x \<in> S"
  9.1877 +  show "eventually (\<lambda>x. ereal (f x) \<in> S) net"
  9.1878 +    using `?r`[THEN topological_tendstoD, OF open_ereal_vimage, OF `open S`]
  9.1879 +    using `ereal x \<in> S` by auto
  9.1880 +qed
  9.1881 +
  9.1882 +lemma lim_real_of_ereal[simp]:
  9.1883 +  assumes lim: "(f ---> ereal x) net"
  9.1884 +  shows "((\<lambda>x. real (f x)) ---> x) net"
  9.1885 +proof (intro topological_tendstoI)
  9.1886 +  fix S assume "open S" "x \<in> S"
  9.1887 +  then have S: "open S" "ereal x \<in> ereal ` S"
  9.1888 +    by (simp_all add: inj_image_mem_iff)
  9.1889 +  have "\<forall>x. f x \<in> ereal ` S \<longrightarrow> real (f x) \<in> S" by auto
  9.1890 +  from this lim[THEN topological_tendstoD, OF open_ereal, OF S]
  9.1891 +  show "eventually (\<lambda>x. real (f x) \<in> S) net"
  9.1892 +    by (rule eventually_mono)
  9.1893 +qed
  9.1894 +
  9.1895 +lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= ereal B)" (is "?l = ?r")
  9.1896 +proof
  9.1897 +  assume ?r
  9.1898 +  show ?l
  9.1899 +    apply(rule topological_tendstoI)
  9.1900 +    unfolding eventually_sequentially
  9.1901 +  proof-
  9.1902 +    fix S :: "ereal set" assume "open S" "\<infinity> : S"
  9.1903 +    from open_PInfty[OF this] guess B .. note B=this
  9.1904 +    from `?r`[rule_format,of "B+1"] guess N .. note N=this
  9.1905 +    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
  9.1906 +    proof safe case goal1
  9.1907 +      have "ereal B < ereal (B + 1)" by auto
  9.1908 +      also have "... <= f n" using goal1 N by auto
  9.1909 +      finally show ?case using B by fastsimp
  9.1910 +    qed
  9.1911 +  qed
  9.1912 +next
  9.1913 +  assume ?l
  9.1914 +  show ?r
  9.1915 +  proof fix B::real have "open {ereal B<..}" "\<infinity> : {ereal B<..}" by auto
  9.1916 +    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
  9.1917 +    guess N .. note N=this
  9.1918 +    show "EX N. ALL n>=N. ereal B <= f n" apply(rule_tac x=N in exI) using N by auto
  9.1919 +  qed
  9.1920 +qed
  9.1921 +
  9.1922 +
  9.1923 +lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= ereal B)" (is "?l = ?r")
  9.1924 +proof
  9.1925 +  assume ?r
  9.1926 +  show ?l
  9.1927 +    apply(rule topological_tendstoI)
  9.1928 +    unfolding eventually_sequentially
  9.1929 +  proof-
  9.1930 +    fix S :: "ereal set"
  9.1931 +    assume "open S" "(-\<infinity>) : S"
  9.1932 +    from open_MInfty[OF this] guess B .. note B=this
  9.1933 +    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
  9.1934 +    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
  9.1935 +    proof safe case goal1
  9.1936 +      have "ereal (B - 1) >= f n" using goal1 N by auto
  9.1937 +      also have "... < ereal B" by auto
  9.1938 +      finally show ?case using B by fastsimp
  9.1939 +    qed
  9.1940 +  qed
  9.1941 +next assume ?l show ?r
  9.1942 +  proof fix B::real have "open {..<ereal B}" "(-\<infinity>) : {..<ereal B}" by auto
  9.1943 +    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
  9.1944 +    guess N .. note N=this
  9.1945 +    show "EX N. ALL n>=N. ereal B >= f n" apply(rule_tac x=N in exI) using N by auto
  9.1946 +  qed
  9.1947 +qed
  9.1948 +
  9.1949 +
  9.1950 +lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= ereal B" shows "l ~= \<infinity>"
  9.1951 +proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
  9.1952 +  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
  9.1953 +  guess N .. note N=this[rule_format,OF le_refl]
  9.1954 +  hence "ereal ?B <= ereal B" using assms(2)[of N] by(rule order_trans)
  9.1955 +  hence "ereal ?B < ereal ?B" apply (rule le_less_trans) by auto
  9.1956 +  thus False by auto
  9.1957 +qed
  9.1958 +
  9.1959 +
  9.1960 +lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= ereal B" shows "l ~= (-\<infinity>)"
  9.1961 +proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
  9.1962 +  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
  9.1963 +  guess N .. note N=this[rule_format,OF le_refl]
  9.1964 +  hence "ereal B <= ereal ?B" using assms(2)[of N] order_trans[of "ereal B" "f N" "ereal(B - 1)"] by blast
  9.1965 +  thus False by auto
  9.1966 +qed
  9.1967 +
  9.1968 +
  9.1969 +lemma tendsto_explicit:
  9.1970 +  "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
  9.1971 +  unfolding tendsto_def eventually_sequentially by auto
  9.1972 +
  9.1973 +
  9.1974 +lemma tendsto_obtains_N:
  9.1975 +  assumes "f ----> f0"
  9.1976 +  assumes "open S" "f0 : S"
  9.1977 +  obtains N where "ALL n>=N. f n : S"
  9.1978 +  using tendsto_explicit[of f f0] assms by auto
  9.1979 +
  9.1980 +
  9.1981 +lemma tail_same_limit:
  9.1982 +  fixes X Y N
  9.1983 +  assumes "X ----> L" "ALL n>=N. X n = Y n"
  9.1984 +  shows "Y ----> L"
  9.1985 +proof-
  9.1986 +{ fix S assume "open S" and "L:S"
  9.1987 +  from this obtain N1 where "ALL n>=N1. X n : S"
  9.1988 +     using assms unfolding tendsto_def eventually_sequentially by auto
  9.1989 +  hence "ALL n>=max N N1. Y n : S" using assms by auto
  9.1990 +  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
  9.1991 +}
  9.1992 +thus ?thesis using tendsto_explicit by auto
  9.1993 +qed
  9.1994 +
  9.1995 +
  9.1996 +lemma Lim_bounded_PInfty2:
  9.1997 +assumes lim:"f ----> l" and "ALL n>=N. f n <= ereal B"
  9.1998 +shows "l ~= \<infinity>"
  9.1999 +proof-
  9.2000 +  def g == "(%n. if n>=N then f n else ereal B)"
  9.2001 +  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
  9.2002 +  moreover have "!!n. g n <= ereal B" using g_def assms by auto
  9.2003 +  ultimately show ?thesis using  Lim_bounded_PInfty by auto
  9.2004 +qed
  9.2005 +
  9.2006 +lemma Lim_bounded_ereal:
  9.2007 +  assumes lim:"f ----> (l :: ereal)"
  9.2008 +  and "ALL n>=M. f n <= C"
  9.2009 +  shows "l<=C"
  9.2010 +proof-
  9.2011 +{ assume "l=(-\<infinity>)" hence ?thesis by auto }
  9.2012 +moreover
  9.2013 +{ assume "~(l=(-\<infinity>))"
  9.2014 +  { assume "C=\<infinity>" hence ?thesis by auto }
  9.2015 +  moreover
  9.2016 +  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
  9.2017 +    hence "l=(-\<infinity>)" using assms
  9.2018 +       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
  9.2019 +    hence ?thesis by auto }
  9.2020 +  moreover
  9.2021 +  { assume "EX B. C = ereal B"
  9.2022 +    from this obtain B where B_def: "C=ereal B" by auto
  9.2023 +    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
  9.2024 +    from this obtain m where m_def: "ereal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
  9.2025 +    from this obtain N where N_def: "ALL n>=N. f n : {ereal(m - 1) <..< ereal(m+1)}"
  9.2026 +       apply (subst tendsto_obtains_N[of f l "{ereal(m - 1) <..< ereal(m+1)}"]) using assms by auto
  9.2027 +    { fix n assume "n>=N"
  9.2028 +      hence "EX r. ereal r = f n" using N_def by (cases "f n") auto
  9.2029 +    } from this obtain g where g_def: "ALL n>=N. ereal (g n) = f n" by metis
  9.2030 +    hence "(%n. ereal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
  9.2031 +    hence *: "(%n. g n) ----> m" using m_def by auto
  9.2032 +    { fix n assume "n>=max N M"
  9.2033 +      hence "ereal (g n) <= ereal B" using assms g_def B_def by auto
  9.2034 +      hence "g n <= B" by auto
  9.2035 +    } hence "EX N. ALL n>=N. g n <= B" by blast
  9.2036 +    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
  9.2037 +    hence ?thesis using m_def B_def by auto
  9.2038 +  } ultimately have ?thesis by (cases C) auto
  9.2039 +} ultimately show ?thesis by blast
  9.2040 +qed
  9.2041 +
  9.2042 +lemma real_of_ereal_mult[simp]:
  9.2043 +  fixes a b :: ereal shows "real (a * b) = real a * real b"
  9.2044 +  by (cases rule: ereal2_cases[of a b]) auto
  9.2045 +
  9.2046 +lemma real_of_ereal_eq_0:
  9.2047 +  fixes x :: ereal shows "real x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity> \<or> x = 0"
  9.2048 +  by (cases x) auto
  9.2049 +
  9.2050 +lemma tendsto_ereal_realD:
  9.2051 +  fixes f :: "'a \<Rightarrow> ereal"
  9.2052 +  assumes "x \<noteq> 0" and tendsto: "((\<lambda>x. ereal (real (f x))) ---> x) net"
  9.2053 +  shows "(f ---> x) net"
  9.2054 +proof (intro topological_tendstoI)
  9.2055 +  fix S assume S: "open S" "x \<in> S"
  9.2056 +  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}" by auto
  9.2057 +  from tendsto[THEN topological_tendstoD, OF this]
  9.2058 +  show "eventually (\<lambda>x. f x \<in> S) net"
  9.2059 +    by (rule eventually_rev_mp) (auto simp: ereal_real real_of_ereal_0)
  9.2060 +qed
  9.2061 +
  9.2062 +lemma tendsto_ereal_realI:
  9.2063 +  fixes f :: "'a \<Rightarrow> ereal"
  9.2064 +  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>" and tendsto: "(f ---> x) net"
  9.2065 +  shows "((\<lambda>x. ereal (real (f x))) ---> x) net"
  9.2066 +proof (intro topological_tendstoI)
  9.2067 +  fix S assume "open S" "x \<in> S"
  9.2068 +  with x have "open (S - {\<infinity>, -\<infinity>})" "x \<in> S - {\<infinity>, -\<infinity>}" by auto
  9.2069 +  from tendsto[THEN topological_tendstoD, OF this]
  9.2070 +  show "eventually (\<lambda>x. ereal (real (f x)) \<in> S) net"
  9.2071 +    by (elim eventually_elim1) (auto simp: ereal_real)
  9.2072 +qed
  9.2073 +
  9.2074 +lemma ereal_mult_cancel_left:
  9.2075 +  fixes a b c :: ereal shows "a * b = a * c \<longleftrightarrow>
  9.2076 +    ((\<bar>a\<bar> = \<infinity> \<and> 0 < b * c) \<or> a = 0 \<or> b = c)"
  9.2077 +  by (cases rule: ereal3_cases[of a b c])
  9.2078 +     (simp_all add: zero_less_mult_iff)
  9.2079 +
  9.2080 +lemma ereal_inj_affinity:
  9.2081 +  fixes m t :: ereal
  9.2082 +  assumes "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" "\<bar>t\<bar> \<noteq> \<infinity>"
  9.2083 +  shows "inj_on (\<lambda>x. m * x + t) A"
  9.2084 +  using assms
  9.2085 +  by (cases rule: ereal2_cases[of m t])
  9.2086 +     (auto intro!: inj_onI simp: ereal_add_cancel_right ereal_mult_cancel_left)
  9.2087 +
  9.2088 +lemma ereal_PInfty_eq_plus[simp]:
  9.2089 +  fixes a b :: ereal
  9.2090 +  shows "\<infinity> = a + b \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
  9.2091 +  by (cases rule: ereal2_cases[of a b]) auto
  9.2092 +
  9.2093 +lemma ereal_MInfty_eq_plus[simp]:
  9.2094 +  fixes a b :: ereal
  9.2095 +  shows "-\<infinity> = a + b \<longleftrightarrow> (a = -\<infinity> \<and> b \<noteq> \<infinity>) \<or> (b = -\<infinity> \<and> a \<noteq> \<infinity>)"
  9.2096 +  by (cases rule: ereal2_cases[of a b]) auto
  9.2097 +
  9.2098 +lemma ereal_less_divide_pos:
  9.2099 +  fixes x y :: ereal
  9.2100 +  shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y < z / x \<longleftrightarrow> x * y < z"
  9.2101 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.2102 +
  9.2103 +lemma ereal_divide_less_pos:
  9.2104 +  fixes x y z :: ereal
  9.2105 +  shows "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y / x < z \<longleftrightarrow> y < x * z"
  9.2106 +  by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps)
  9.2107 +
  9.2108 +lemma ereal_divide_eq:
  9.2109 +  fixes a b c :: ereal
  9.2110 +  shows "b \<noteq> 0 \<Longrightarrow> \<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> a / b = c \<longleftrightarrow> a = b * c"
  9.2111 +  by (cases rule: ereal3_cases[of a b c])
  9.2112 +     (simp_all add: field_simps)
  9.2113 +
  9.2114 +lemma ereal_inverse_not_MInfty[simp]: "inverse (a::ereal) \<noteq> -\<infinity>"
  9.2115 +  by (cases a) auto
  9.2116 +
  9.2117 +lemma ereal_mult_m1[simp]: "x * ereal (-1) = -x"
  9.2118 +  by (cases x) auto
  9.2119 +
  9.2120 +lemma ereal_LimI_finite:
  9.2121 +  fixes x :: ereal
  9.2122 +  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
  9.2123 +  assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r"
  9.2124 +  shows "u ----> x"
  9.2125 +proof (rule topological_tendstoI, unfold eventually_sequentially)
  9.2126 +  obtain rx where rx_def: "x=ereal rx" using assms by (cases x) auto
  9.2127 +  fix S assume "open S" "x : S"
  9.2128 +  then have "open (ereal -` S)" unfolding open_ereal_def by auto
  9.2129 +  with `x \<in> S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> ereal y \<in> S"
  9.2130 +    unfolding open_real_def rx_def by auto
  9.2131 +  then obtain n where
  9.2132 +    upper: "!!N. n <= N ==> u N < x + ereal r" and
  9.2133 +    lower: "!!N. n <= N ==> x < u N + ereal r" using assms(2)[of "ereal r"] by auto
  9.2134 +  show "EX N. ALL n>=N. u n : S"
  9.2135 +  proof (safe intro!: exI[of _ n])
  9.2136 +    fix N assume "n <= N"
  9.2137 +    from upper[OF this] lower[OF this] assms `0 < r`
  9.2138 +    have "u N ~: {\<infinity>,(-\<infinity>)}" by auto
  9.2139 +    from this obtain ra where ra_def: "(u N) = ereal ra" by (cases "u N") auto
  9.2140 +    hence "rx < ra + r" and "ra < rx + r"
  9.2141 +       using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto
  9.2142 +    hence "dist (real (u N)) rx < r"
  9.2143 +      using rx_def ra_def
  9.2144 +      by (auto simp: dist_real_def abs_diff_less_iff field_simps)
  9.2145 +    from dist[OF this] show "u N : S" using `u N  ~: {\<infinity>, -\<infinity>}`
  9.2146 +      by (auto simp: ereal_real split: split_if_asm)
  9.2147 +  qed
  9.2148 +qed
  9.2149 +
  9.2150 +lemma ereal_LimI_finite_iff:
  9.2151 +  fixes x :: ereal
  9.2152 +  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
  9.2153 +  shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))"
  9.2154 +  (is "?lhs <-> ?rhs")
  9.2155 +proof
  9.2156 +  assume lim: "u ----> x"
  9.2157 +  { fix r assume "(r::ereal)>0"
  9.2158 +    from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}"
  9.2159 +       apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
  9.2160 +       using lim ereal_between[of x r] assms `r>0` by auto
  9.2161 +    hence "EX N. ALL n>=N. u n < x + r & x < u n + r"
  9.2162 +      using ereal_minus_less[of r x] by (cases r) auto
  9.2163 +  } then show "?rhs" by auto
  9.2164 +next
  9.2165 +  assume ?rhs then show "u ----> x"
  9.2166 +    using ereal_LimI_finite[of x] assms by auto
  9.2167 +qed
  9.2168 +
  9.2169 +
  9.2170 +subsubsection {* @{text Liminf} and @{text Limsup} *}
  9.2171 +
  9.2172 +definition
  9.2173 +  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
  9.2174 +
  9.2175 +definition
  9.2176 +  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
  9.2177 +
  9.2178 +lemma Liminf_Sup:
  9.2179 +  fixes f :: "'a => 'b::{complete_lattice, linorder}"
  9.2180 +  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
  9.2181 +  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
  9.2182 +
  9.2183 +lemma Limsup_Inf:
  9.2184 +  fixes f :: "'a => 'b::{complete_lattice, linorder}"
  9.2185 +  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
  9.2186 +  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
  9.2187 +
  9.2188 +lemma ereal_SupI:
  9.2189 +  fixes x :: ereal
  9.2190 +  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
  9.2191 +  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
  9.2192 +  shows "Sup A = x"
  9.2193 +  unfolding Sup_ereal_def
  9.2194 +  using assms by (auto intro!: Least_equality)
  9.2195 +
  9.2196 +lemma ereal_InfI:
  9.2197 +  fixes x :: ereal
  9.2198 +  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
  9.2199 +  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
  9.2200 +  shows "Inf A = x"
  9.2201 +  unfolding Inf_ereal_def
  9.2202 +  using assms by (auto intro!: Greatest_equality)
  9.2203 +
  9.2204 +lemma Limsup_const:
  9.2205 +  fixes c :: "'a::{complete_lattice, linorder}"
  9.2206 +  assumes ntriv: "\<not> trivial_limit net"
  9.2207 +  shows "Limsup net (\<lambda>x. c) = c"
  9.2208 +  unfolding Limsup_Inf
  9.2209 +proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
  9.2210 +  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
  9.2211 +  show "c \<le> x"
  9.2212 +  proof (rule ccontr)
  9.2213 +    assume "\<not> c \<le> x" then have "x < c" by auto
  9.2214 +    then show False using ntriv * by (auto simp: trivial_limit_def)
  9.2215 +  qed
  9.2216 +qed auto
  9.2217 +
  9.2218 +lemma Liminf_const:
  9.2219 +  fixes c :: "'a::{complete_lattice, linorder}"
  9.2220 +  assumes ntriv: "\<not> trivial_limit net"
  9.2221 +  shows "Liminf net (\<lambda>x. c) = c"
  9.2222 +  unfolding Liminf_Sup
  9.2223 +proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
  9.2224 +  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
  9.2225 +  show "x \<le> c"
  9.2226 +  proof (rule ccontr)
  9.2227 +    assume "\<not> x \<le> c" then have "c < x" by auto
  9.2228 +    then show False using ntriv * by (auto simp: trivial_limit_def)
  9.2229 +  qed
  9.2230 +qed auto
  9.2231 +
  9.2232 +lemma mono_set:
  9.2233 +  fixes S :: "('a::order) set"
  9.2234 +  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
  9.2235 +  by (auto simp: mono_def mem_def)
  9.2236 +
  9.2237 +lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
  9.2238 +lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
  9.2239 +lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
  9.2240 +lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
  9.2241 +
  9.2242 +lemma mono_set_iff:
  9.2243 +  fixes S :: "'a::{linorder,complete_lattice} set"
  9.2244 +  defines "a \<equiv> Inf S"
  9.2245 +  shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
  9.2246 +proof
  9.2247 +  assume "mono S"
  9.2248 +  then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set)
  9.2249 +  show ?c
  9.2250 +  proof cases
  9.2251 +    assume "a \<in> S"
  9.2252 +    show ?c
  9.2253 +      using mono[OF _ `a \<in> S`]
  9.2254 +      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
  9.2255 +  next
  9.2256 +    assume "a \<notin> S"
  9.2257 +    have "S = {a <..}"
  9.2258 +    proof safe
  9.2259 +      fix x assume "x \<in> S"
  9.2260 +      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
  9.2261 +      then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
  9.2262 +    next
  9.2263 +      fix x assume "a < x"
  9.2264 +      then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff ..
  9.2265 +      with mono[of y x] show "x \<in> S" by auto
  9.2266 +    qed
  9.2267 +    then show ?c ..
  9.2268 +  qed
  9.2269 +qed auto
  9.2270 +
  9.2271 +lemma lim_imp_Liminf:
  9.2272 +  fixes f :: "'a \<Rightarrow> ereal"
  9.2273 +  assumes ntriv: "\<not> trivial_limit net"
  9.2274 +  assumes lim: "(f ---> f0) net"
  9.2275 +  shows "Liminf net f = f0"
  9.2276 +  unfolding Liminf_Sup
  9.2277 +proof (safe intro!: ereal_SupI)
  9.2278 +  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
  9.2279 +  show "y \<le> f0"
  9.2280 +  proof (rule ereal_le_ereal)
  9.2281 +    fix B assume "B < y"
  9.2282 +    { assume "f0 < B"
  9.2283 +      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
  9.2284 +         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
  9.2285 +         by (auto intro: eventually_conj)
  9.2286 +      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
  9.2287 +      finally have False using ntriv[unfolded trivial_limit_def] by auto
  9.2288 +    } then show "B \<le> f0" by (metis linorder_le_less_linear)
  9.2289 +  qed
  9.2290 +next
  9.2291 +  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
  9.2292 +  show "f0 \<le> y"
  9.2293 +  proof (safe intro!: *[rule_format])
  9.2294 +    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
  9.2295 +      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
  9.2296 +  qed
  9.2297 +qed
  9.2298 +
  9.2299 +lemma ereal_Liminf_le_Limsup:
  9.2300 +  fixes f :: "'a \<Rightarrow> ereal"
  9.2301 +  assumes ntriv: "\<not> trivial_limit net"
  9.2302 +  shows "Liminf net f \<le> Limsup net f"
  9.2303 +  unfolding Limsup_Inf Liminf_Sup
  9.2304 +proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
  9.2305 +  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
  9.2306 +  show "u \<le> v"
  9.2307 +  proof (rule ccontr)
  9.2308 +    assume "\<not> u \<le> v"
  9.2309 +    then obtain t where "t < u" "v < t"
  9.2310 +      using ereal_dense[of v u] by (auto simp: not_le)
  9.2311 +    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
  9.2312 +      using * by (auto intro: eventually_conj)
  9.2313 +    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
  9.2314 +    finally show False using ntriv by (auto simp: trivial_limit_def)
  9.2315 +  qed
  9.2316 +qed
  9.2317 +
  9.2318 +lemma Liminf_mono:
  9.2319 +  fixes f g :: "'a => ereal"
  9.2320 +  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
  9.2321 +  shows "Liminf net f \<le> Liminf net g"
  9.2322 +  unfolding Liminf_Sup
  9.2323 +proof (safe intro!: Sup_mono bexI)
  9.2324 +  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
  9.2325 +  then have "eventually (\<lambda>x. y < f x) net" by auto
  9.2326 +  then show "eventually (\<lambda>x. y < g x) net"
  9.2327 +    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
  9.2328 +qed simp
  9.2329 +
  9.2330 +lemma Liminf_eq:
  9.2331 +  fixes f g :: "'a \<Rightarrow> ereal"
  9.2332 +  assumes "eventually (\<lambda>x. f x = g x) net"
  9.2333 +  shows "Liminf net f = Liminf net g"
  9.2334 +  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
  9.2335 +
  9.2336 +lemma Liminf_mono_all:
  9.2337 +  fixes f g :: "'a \<Rightarrow> ereal"
  9.2338 +  assumes "\<And>x. f x \<le> g x"
  9.2339 +  shows "Liminf net f \<le> Liminf net g"
  9.2340 +  using assms by (intro Liminf_mono always_eventually) auto
  9.2341 +
  9.2342 +lemma Limsup_mono:
  9.2343 +  fixes f g :: "'a \<Rightarrow> ereal"
  9.2344 +  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
  9.2345 +  shows "Limsup net f \<le> Limsup net g"
  9.2346 +  unfolding Limsup_Inf
  9.2347 +proof (safe intro!: Inf_mono bexI)
  9.2348 +  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
  9.2349 +  then have "eventually (\<lambda>x. g x < y) net" by auto
  9.2350 +  then show "eventually (\<lambda>x. f x < y) net"
  9.2351 +    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
  9.2352 +qed simp
  9.2353 +
  9.2354 +lemma Limsup_mono_all:
  9.2355 +  fixes f g :: "'a \<Rightarrow> ereal"
  9.2356 +  assumes "\<And>x. f x \<le> g x"
  9.2357 +  shows "Limsup net f \<le> Limsup net g"
  9.2358 +  using assms by (intro Limsup_mono always_eventually) auto
  9.2359 +
  9.2360 +lemma Limsup_eq:
  9.2361 +  fixes f g :: "'a \<Rightarrow> ereal"
  9.2362 +  assumes "eventually (\<lambda>x. f x = g x) net"
  9.2363 +  shows "Limsup net f = Limsup net g"
  9.2364 +  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
  9.2365 +
  9.2366 +abbreviation "liminf \<equiv> Liminf sequentially"
  9.2367 +
  9.2368 +abbreviation "limsup \<equiv> Limsup sequentially"
  9.2369 +
  9.2370 +lemma (in complete_lattice) less_INFD:
  9.2371 +  assumes "y < INFI A f"" i \<in> A" shows "y < f i"
  9.2372 +proof -
  9.2373 +  note `y < INFI A f`
  9.2374 +  also have "INFI A f \<le> f i" using `i \<in> A` by (rule INF_leI)
  9.2375 +  finally show "y < f i" .
  9.2376 +qed
  9.2377 +
  9.2378 +lemma liminf_SUPR_INFI:
  9.2379 +  fixes f :: "nat \<Rightarrow> ereal"
  9.2380 +  shows "liminf f = (SUP n. INF m:{n..}. f m)"
  9.2381 +  unfolding Liminf_Sup eventually_sequentially
  9.2382 +proof (safe intro!: antisym complete_lattice_class.Sup_least)
  9.2383 +  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)"
  9.2384 +  proof (rule ereal_le_ereal)
  9.2385 +    fix y assume "y < x"
  9.2386 +    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
  9.2387 +    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
  9.2388 +    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto
  9.2389 +    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
  9.2390 +  qed
  9.2391 +next
  9.2392 +  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
  9.2393 +  proof (unfold SUPR_def, safe intro!: Sup_mono bexI)
  9.2394 +    fix y n assume "y < INFI {n..} f"
  9.2395 +    from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
  9.2396 +  qed (rule order_refl)
  9.2397 +qed
  9.2398 +
  9.2399 +lemma tail_same_limsup:
  9.2400 +  fixes X Y :: "nat => ereal"
  9.2401 +  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
  9.2402 +  shows "limsup X = limsup Y"
  9.2403 +  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
  9.2404 +
  9.2405 +lemma tail_same_liminf:
  9.2406 +  fixes X Y :: "nat => ereal"
  9.2407 +  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
  9.2408 +  shows "liminf X = liminf Y"
  9.2409 +  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
  9.2410 +
  9.2411 +lemma liminf_mono:
  9.2412 +  fixes X Y :: "nat \<Rightarrow> ereal"
  9.2413 +  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
  9.2414 +  shows "liminf X \<le> liminf Y"
  9.2415 +  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
  9.2416 +
  9.2417 +lemma limsup_mono:
  9.2418 +  fixes X Y :: "nat => ereal"
  9.2419 +  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
  9.2420 +  shows "limsup X \<le> limsup Y"
  9.2421 +  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
  9.2422 +
  9.2423 +declare trivial_limit_sequentially[simp]
  9.2424 +
  9.2425 +lemma
  9.2426 +  fixes X :: "nat \<Rightarrow> ereal"
  9.2427 +  shows ereal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
  9.2428 +    and ereal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
  9.2429 +  unfolding incseq_def decseq_def by auto
  9.2430 +
  9.2431 +lemma liminf_bounded:
  9.2432 +  fixes X Y :: "nat \<Rightarrow> ereal"
  9.2433 +  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
  9.2434 +  shows "C \<le> liminf X"
  9.2435 +  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
  9.2436 +
  9.2437 +lemma limsup_bounded:
  9.2438 +  fixes X Y :: "nat => ereal"
  9.2439 +  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
  9.2440 +  shows "limsup X \<le> C"
  9.2441 +  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
  9.2442 +
  9.2443 +lemma liminf_bounded_iff:
  9.2444 +  fixes x :: "nat \<Rightarrow> ereal"
  9.2445 +  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
  9.2446 +proof safe
  9.2447 +  fix B assume "B < C" "C \<le> liminf x"
  9.2448 +  then have "B < liminf x" by auto
  9.2449 +  then obtain N where "B < (INF m:{N..}. x m)"
  9.2450 +    unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto
  9.2451 +  from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
  9.2452 +next
  9.2453 +  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
  9.2454 +  { fix B assume "B<C"
  9.2455 +    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
  9.2456 +    hence "B \<le> (INF m:{N..}. x m)" by (intro le_INFI) auto
  9.2457 +    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp
  9.2458 +    finally have "B \<le> liminf x" .
  9.2459 +  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
  9.2460 +qed
  9.2461 +
  9.2462 +lemma liminf_subseq_mono:
  9.2463 +  fixes X :: "nat \<Rightarrow> ereal"
  9.2464 +  assumes "subseq r"
  9.2465 +  shows "liminf X \<le> liminf (X \<circ> r) "
  9.2466 +proof-
  9.2467 +  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
  9.2468 +  proof (safe intro!: INF_mono)
  9.2469 +    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
  9.2470 +      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
  9.2471 +  qed
  9.2472 +  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
  9.2473 +qed
  9.2474 +
  9.2475 +lemma ereal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "ereal (real x) = x"
  9.2476 +  using assms by auto
  9.2477 +
  9.2478 +lemma ereal_le_ereal_bounded:
  9.2479 +  fixes x y z :: ereal
  9.2480 +  assumes "z \<le> y"
  9.2481 +  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
  9.2482 +  shows "x \<le> y"
  9.2483 +proof (rule ereal_le_ereal)
  9.2484 +  fix B assume "B < x"
  9.2485 +  show "B \<le> y"
  9.2486 +  proof cases
  9.2487 +    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
  9.2488 +  next
  9.2489 +    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
  9.2490 +  qed
  9.2491 +qed
  9.2492 +
  9.2493 +lemma fixes x y :: ereal
  9.2494 +  shows Sup_atMost[simp]: "Sup {.. y} = y"
  9.2495 +    and Sup_lessThan[simp]: "Sup {..< y} = y"
  9.2496 +    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
  9.2497 +    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
  9.2498 +    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
  9.2499 +  by (auto simp: Sup_ereal_def intro!: Least_equality
  9.2500 +           intro: ereal_le_ereal ereal_le_ereal_bounded[of x])
  9.2501 +
  9.2502 +lemma Sup_greaterThanlessThan[simp]:
  9.2503 +  fixes x y :: ereal assumes "x < y" shows "Sup { x <..< y} = y"
  9.2504 +  unfolding Sup_ereal_def
  9.2505 +proof (intro Least_equality ereal_le_ereal_bounded[of _ _ y])
  9.2506 +  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
  9.2507 +  from ereal_dense[OF `x < y`] guess w .. note w = this
  9.2508 +  with z[THEN bspec, of w] show "x \<le> z" by auto
  9.2509 +qed auto
  9.2510 +
  9.2511 +lemma real_ereal_id: "real o ereal = id"
  9.2512 +proof-
  9.2513 +{ fix x have "(real o ereal) x = id x" by auto }
  9.2514 +from this show ?thesis using ext by blast
  9.2515 +qed
  9.2516 +
  9.2517 +lemma open_image_ereal: "open(UNIV-{ \<infinity> , (-\<infinity> :: ereal)})"
  9.2518 +by (metis range_ereal open_ereal open_UNIV)
  9.2519 +
  9.2520 +lemma ereal_le_distrib:
  9.2521 +  fixes a b c :: ereal shows "c * (a + b) \<le> c * a + c * b"
  9.2522 +  by (cases rule: ereal3_cases[of a b c])
  9.2523 +     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
  9.2524 +
  9.2525 +lemma ereal_pos_distrib:
  9.2526 +  fixes a b c :: ereal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
  9.2527 +  using assms by (cases rule: ereal3_cases[of a b c])
  9.2528 +                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
  9.2529 +
  9.2530 +lemma ereal_pos_le_distrib:
  9.2531 +fixes a b c :: ereal
  9.2532 +assumes "c>=0"
  9.2533 +shows "c * (a + b) <= c * a + c * b"
  9.2534 +  using assms by (cases rule: ereal3_cases[of a b c])
  9.2535 +                 (auto simp add: field_simps)
  9.2536 +
  9.2537 +lemma ereal_max_mono:
  9.2538 +  "[| (a::ereal) <= b; c <= d |] ==> max a c <= max b d"
  9.2539 +  by (metis sup_ereal_def sup_mono)
  9.2540 +
  9.2541 +
  9.2542 +lemma ereal_max_least:
  9.2543 +  "[| (a::ereal) <= x; c <= x |] ==> max a c <= x"
  9.2544 +  by (metis sup_ereal_def sup_least)
  9.2545 +
  9.2546 +end
    10.1 --- a/src/HOL/Library/Extended_Reals.thy	Wed Jul 20 10:11:08 2011 +0200
    10.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    10.3 @@ -1,2535 +0,0 @@
    10.4 -(*  Title:      HOL/Library/Extended_Reals.thy
    10.5 -    Author:     Johannes Hölzl, TU München
    10.6 -    Author:     Robert Himmelmann, TU München
    10.7 -    Author:     Armin Heller, TU München
    10.8 -    Author:     Bogdan Grechuk, University of Edinburgh
    10.9 -*)
   10.10 -
   10.11 -header {* Extended real number line *}
   10.12 -
   10.13 -theory Extended_Reals
   10.14 -  imports Complex_Main
   10.15 -begin
   10.16 -
   10.17 -text {*
   10.18 -
   10.19 -For more lemmas about the extended real numbers go to
   10.20 -  @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"}
   10.21 -
   10.22 -*}
   10.23 -
   10.24 -lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
   10.25 -proof
   10.26 -  assume "{x..} = UNIV"
   10.27 -  show "x = bot"
   10.28 -  proof (rule ccontr)
   10.29 -    assume "x \<noteq> bot" then have "bot \<notin> {x..}" by (simp add: le_less)
   10.30 -    then show False using `{x..} = UNIV` by simp
   10.31 -  qed
   10.32 -qed auto
   10.33 -
   10.34 -lemma SUPR_pair:
   10.35 -  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
   10.36 -  by (rule antisym) (auto intro!: SUP_leI le_SUPI2)
   10.37 -
   10.38 -lemma INFI_pair:
   10.39 -  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
   10.40 -  by (rule antisym) (auto intro!: le_INFI INF_leI2)
   10.41 -
   10.42 -subsection {* Definition and basic properties *}
   10.43 -
   10.44 -datatype extreal = extreal real | PInfty | MInfty
   10.45 -
   10.46 -notation (xsymbols)
   10.47 -  PInfty  ("\<infinity>")
   10.48 -
   10.49 -notation (HTML output)
   10.50 -  PInfty  ("\<infinity>")
   10.51 -
   10.52 -declare [[coercion "extreal :: real \<Rightarrow> extreal"]]
   10.53 -
   10.54 -instantiation extreal :: uminus
   10.55 -begin
   10.56 -  fun uminus_extreal where
   10.57 -    "- (extreal r) = extreal (- r)"
   10.58 -  | "- \<infinity> = MInfty"
   10.59 -  | "- MInfty = \<infinity>"
   10.60 -  instance ..
   10.61 -end
   10.62 -
   10.63 -lemma inj_extreal[simp]: "inj_on extreal A"
   10.64 -  unfolding inj_on_def by auto
   10.65 -
   10.66 -lemma MInfty_neq_PInfty[simp]:
   10.67 -  "\<infinity> \<noteq> - \<infinity>" "- \<infinity> \<noteq> \<infinity>" by simp_all
   10.68 -
   10.69 -lemma MInfty_neq_extreal[simp]:
   10.70 -  "extreal r \<noteq> - \<infinity>" "- \<infinity> \<noteq> extreal r" by simp_all
   10.71 -
   10.72 -lemma MInfinity_cases[simp]:
   10.73 -  "(case - \<infinity> of extreal r \<Rightarrow> f r | \<infinity> \<Rightarrow> y | MInfinity \<Rightarrow> z) = z"
   10.74 -  by simp
   10.75 -
   10.76 -lemma extreal_uminus_uminus[simp]:
   10.77 -  fixes a :: extreal shows "- (- a) = a"
   10.78 -  by (cases a) simp_all
   10.79 -
   10.80 -lemma MInfty_eq[simp, code_post]:
   10.81 -  "MInfty = - \<infinity>" by simp
   10.82 -
   10.83 -declare uminus_extreal.simps(2)[code_inline, simp del]
   10.84 -
   10.85 -lemma extreal_cases[case_names real PInf MInf, cases type: extreal]:
   10.86 -  assumes "\<And>r. x = extreal r \<Longrightarrow> P"
   10.87 -  assumes "x = \<infinity> \<Longrightarrow> P"
   10.88 -  assumes "x = -\<infinity> \<Longrightarrow> P"
   10.89 -  shows P
   10.90 -  using assms by (cases x) auto
   10.91 -
   10.92 -lemmas extreal2_cases = extreal_cases[case_product extreal_cases]
   10.93 -lemmas extreal3_cases = extreal2_cases[case_product extreal_cases]
   10.94 -
   10.95 -lemma extreal_uminus_eq_iff[simp]:
   10.96 -  fixes a b :: extreal shows "-a = -b \<longleftrightarrow> a = b"
   10.97 -  by (cases rule: extreal2_cases[of a b]) simp_all
   10.98 -
   10.99 -function of_extreal :: "extreal \<Rightarrow> real" where
  10.100 -"of_extreal (extreal r) = r" |
  10.101 -"of_extreal \<infinity> = 0" |
  10.102 -"of_extreal (-\<infinity>) = 0"
  10.103 -  by (auto intro: extreal_cases)
  10.104 -termination proof qed (rule wf_empty)
  10.105 -
  10.106 -defs (overloaded)
  10.107 -  real_of_extreal_def [code_unfold]: "real \<equiv> of_extreal"
  10.108 -
  10.109 -lemma real_of_extreal[simp]:
  10.110 -    "real (- x :: extreal) = - (real x)"
  10.111 -    "real (extreal r) = r"
  10.112 -    "real \<infinity> = 0"
  10.113 -  by (cases x) (simp_all add: real_of_extreal_def)
  10.114 -
  10.115 -lemma range_extreal[simp]: "range extreal = UNIV - {\<infinity>, -\<infinity>}"
  10.116 -proof safe
  10.117 -  fix x assume "x \<notin> range extreal" "x \<noteq> \<infinity>"
  10.118 -  then show "x = -\<infinity>" by (cases x) auto
  10.119 -qed auto
  10.120 -
  10.121 -lemma extreal_range_uminus[simp]: "range uminus = (UNIV::extreal set)"
  10.122 -proof safe
  10.123 -  fix x :: extreal show "x \<in> range uminus" by (intro image_eqI[of _ _ "-x"]) auto
  10.124 -qed auto
  10.125 -
  10.126 -instantiation extreal :: number
  10.127 -begin
  10.128 -definition [simp]: "number_of x = extreal (number_of x)"
  10.129 -instance proof qed
  10.130 -end
  10.131 -
  10.132 -instantiation extreal :: abs
  10.133 -begin
  10.134 -  function abs_extreal where
  10.135 -    "\<bar>extreal r\<bar> = extreal \<bar>r\<bar>"
  10.136 -  | "\<bar>-\<infinity>\<bar> = \<infinity>"
  10.137 -  | "\<bar>\<infinity>\<bar> = \<infinity>"
  10.138 -  by (auto intro: extreal_cases)
  10.139 -  termination proof qed (rule wf_empty)
  10.140 -  instance ..
  10.141 -end
  10.142 -
  10.143 -lemma abs_eq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> = \<infinity> ; x = \<infinity> \<Longrightarrow> P ; x = -\<infinity> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
  10.144 -  by (cases x) auto
  10.145 -
  10.146 -lemma abs_neq_infinity_cases[elim!]: "\<lbrakk> \<bar>x\<bar> \<noteq> \<infinity> ; \<And>r. x = extreal r \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
  10.147 -  by (cases x) auto
  10.148 -
  10.149 -lemma abs_extreal_uminus[simp]: "\<bar>- x\<bar> = \<bar>x::extreal\<bar>"
  10.150 -  by (cases x) auto
  10.151 -
  10.152 -subsubsection "Addition"
  10.153 -
  10.154 -instantiation extreal :: comm_monoid_add
  10.155 -begin
  10.156 -
  10.157 -definition "0 = extreal 0"
  10.158 -
  10.159 -function plus_extreal where
  10.160 -"extreal r + extreal p = extreal (r + p)" |
  10.161 -"\<infinity> + a = \<infinity>" |
  10.162 -"a + \<infinity> = \<infinity>" |
  10.163 -"extreal r + -\<infinity> = - \<infinity>" |
  10.164 -"-\<infinity> + extreal p = -\<infinity>" |
  10.165 -"-\<infinity> + -\<infinity> = -\<infinity>"
  10.166 -proof -
  10.167 -  case (goal1 P x)
  10.168 -  moreover then obtain a b where "x = (a, b)" by (cases x) auto
  10.169 -  ultimately show P
  10.170 -   by (cases rule: extreal2_cases[of a b]) auto
  10.171 -qed auto
  10.172 -termination proof qed (rule wf_empty)
  10.173 -
  10.174 -lemma Infty_neq_0[simp]:
  10.175 -  "\<infinity> \<noteq> 0" "0 \<noteq> \<infinity>"
  10.176 -  "-\<infinity> \<noteq> 0" "0 \<noteq> -\<infinity>"
  10.177 -  by (simp_all add: zero_extreal_def)
  10.178 -
  10.179 -lemma extreal_eq_0[simp]:
  10.180 -  "extreal r = 0 \<longleftrightarrow> r = 0"
  10.181 -  "0 = extreal r \<longleftrightarrow> r = 0"
  10.182 -  unfolding zero_extreal_def by simp_all
  10.183 -
  10.184 -instance
  10.185 -proof
  10.186 -  fix a :: extreal show "0 + a = a"
  10.187 -    by (cases a) (simp_all add: zero_extreal_def)
  10.188 -  fix b :: extreal show "a + b = b + a"
  10.189 -    by (cases rule: extreal2_cases[of a b]) simp_all
  10.190 -  fix c :: extreal show "a + b + c = a + (b + c)"
  10.191 -    by (cases rule: extreal3_cases[of a b c]) simp_all
  10.192 -qed
  10.193 -end
  10.194 -
  10.195 -lemma real_of_extreal_0[simp]: "real (0::extreal) = 0"
  10.196 -  unfolding real_of_extreal_def zero_extreal_def by simp
  10.197 -
  10.198 -lemma abs_extreal_zero[simp]: "\<bar>0\<bar> = (0::extreal)"
  10.199 -  unfolding zero_extreal_def abs_extreal.simps by simp
  10.200 -
  10.201 -lemma extreal_uminus_zero[simp]:
  10.202 -  "- 0 = (0::extreal)"
  10.203 -  by (simp add: zero_extreal_def)
  10.204 -
  10.205 -lemma extreal_uminus_zero_iff[simp]:
  10.206 -  fixes a :: extreal shows "-a = 0 \<longleftrightarrow> a = 0"
  10.207 -  by (cases a) simp_all
  10.208 -
  10.209 -lemma extreal_plus_eq_PInfty[simp]:
  10.210 -  shows "a + b = \<infinity> \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
  10.211 -  by (cases rule: extreal2_cases[of a b]) auto
  10.212 -
  10.213 -lemma extreal_plus_eq_MInfty[simp]:
  10.214 -  shows "a + b = -\<infinity> \<longleftrightarrow>
  10.215 -    (a = -\<infinity> \<or> b = -\<infinity>) \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>"
  10.216 -  by (cases rule: extreal2_cases[of a b]) auto
  10.217 -
  10.218 -lemma extreal_add_cancel_left:
  10.219 -  assumes "a \<noteq> -\<infinity>"
  10.220 -  shows "a + b = a + c \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
  10.221 -  using assms by (cases rule: extreal3_cases[of a b c]) auto
  10.222 -
  10.223 -lemma extreal_add_cancel_right:
  10.224 -  assumes "a \<noteq> -\<infinity>"
  10.225 -  shows "b + a = c + a \<longleftrightarrow> (a = \<infinity> \<or> b = c)"
  10.226 -  using assms by (cases rule: extreal3_cases[of a b c]) auto
  10.227 -
  10.228 -lemma extreal_real:
  10.229 -  "extreal (real x) = (if \<bar>x\<bar> = \<infinity> then 0 else x)"
  10.230 -  by (cases x) simp_all
  10.231 -
  10.232 -lemma real_of_extreal_add:
  10.233 -  fixes a b :: extreal
  10.234 -  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)"
  10.235 -  by (cases rule: extreal2_cases[of a b]) auto
  10.236 -
  10.237 -subsubsection "Linear order on @{typ extreal}"
  10.238 -
  10.239 -instantiation extreal :: linorder
  10.240 -begin
  10.241 -
  10.242 -function less_extreal where
  10.243 -"extreal x < extreal y \<longleftrightarrow> x < y" |
  10.244 -"        \<infinity> < a         \<longleftrightarrow> False" |
  10.245 -"        a < -\<infinity>        \<longleftrightarrow> False" |
  10.246 -"extreal x < \<infinity>         \<longleftrightarrow> True" |
  10.247 -"       -\<infinity> < extreal r \<longleftrightarrow> True" |
  10.248 -"       -\<infinity> < \<infinity>         \<longleftrightarrow> True"
  10.249 -proof -
  10.250 -  case (goal1 P x)
  10.251 -  moreover then obtain a b where "x = (a,b)" by (cases x) auto
  10.252 -  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
  10.253 -qed simp_all
  10.254 -termination by (relation "{}") simp
  10.255 -
  10.256 -definition "x \<le> (y::extreal) \<longleftrightarrow> x < y \<or> x = y"
  10.257 -
  10.258 -lemma extreal_infty_less[simp]:
  10.259 -  "x < \<infinity> \<longleftrightarrow> (x \<noteq> \<infinity>)"
  10.260 -  "-\<infinity> < x \<longleftrightarrow> (x \<noteq> -\<infinity>)"
  10.261 -  by (cases x, simp_all) (cases x, simp_all)
  10.262 -
  10.263 -lemma extreal_infty_less_eq[simp]:
  10.264 -  "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
  10.265 -  "x \<le> -\<infinity> \<longleftrightarrow> x = -\<infinity>"
  10.266 -  by (auto simp add: less_eq_extreal_def)
  10.267 -
  10.268 -lemma extreal_less[simp]:
  10.269 -  "extreal r < 0 \<longleftrightarrow> (r < 0)"
  10.270 -  "0 < extreal r \<longleftrightarrow> (0 < r)"
  10.271 -  "0 < \<infinity>"
  10.272 -  "-\<infinity> < 0"
  10.273 -  by (simp_all add: zero_extreal_def)
  10.274 -
  10.275 -lemma extreal_less_eq[simp]:
  10.276 -  "x \<le> \<infinity>"
  10.277 -  "-\<infinity> \<le> x"
  10.278 -  "extreal r \<le> extreal p \<longleftrightarrow> r \<le> p"
  10.279 -  "extreal r \<le> 0 \<longleftrightarrow> r \<le> 0"
  10.280 -  "0 \<le> extreal r \<longleftrightarrow> 0 \<le> r"
  10.281 -  by (auto simp add: less_eq_extreal_def zero_extreal_def)
  10.282 -
  10.283 -lemma extreal_infty_less_eq2:
  10.284 -  "a \<le> b \<Longrightarrow> a = \<infinity> \<Longrightarrow> b = \<infinity>"
  10.285 -  "a \<le> b \<Longrightarrow> b = -\<infinity> \<Longrightarrow> a = -\<infinity>"
  10.286 -  by simp_all
  10.287 -
  10.288 -instance
  10.289 -proof
  10.290 -  fix x :: extreal show "x \<le> x"
  10.291 -    by (cases x) simp_all
  10.292 -  fix y :: extreal show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
  10.293 -    by (cases rule: extreal2_cases[of x y]) auto
  10.294 -  show "x \<le> y \<or> y \<le> x "
  10.295 -    by (cases rule: extreal2_cases[of x y]) auto
  10.296 -  { assume "x \<le> y" "y \<le> x" then show "x = y"
  10.297 -    by (cases rule: extreal2_cases[of x y]) auto }
  10.298 -  { fix z assume "x \<le> y" "y \<le> z" then show "x \<le> z"
  10.299 -    by (cases rule: extreal3_cases[of x y z]) auto }
  10.300 -qed
  10.301 -end
  10.302 -
  10.303 -instance extreal :: ordered_ab_semigroup_add
  10.304 -proof
  10.305 -  fix a b c :: extreal assume "a \<le> b" then show "c + a \<le> c + b"
  10.306 -    by (cases rule: extreal3_cases[of a b c]) auto
  10.307 -qed
  10.308 -
  10.309 -lemma real_of_extreal_positive_mono:
  10.310 -  "\<lbrakk>0 \<le> x; x \<le> y; y \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> real x \<le> real y"
  10.311 -  by (cases rule: extreal2_cases[of x y]) auto
  10.312 -
  10.313 -lemma extreal_MInfty_lessI[intro, simp]:
  10.314 -  "a \<noteq> -\<infinity> \<Longrightarrow> -\<infinity> < a"
  10.315 -  by (cases a) auto
  10.316 -
  10.317 -lemma extreal_less_PInfty[intro, simp]:
  10.318 -  "a \<noteq> \<infinity> \<Longrightarrow> a < \<infinity>"
  10.319 -  by (cases a) auto
  10.320 -
  10.321 -lemma extreal_less_extreal_Ex:
  10.322 -  fixes a b :: extreal
  10.323 -  shows "x < extreal r \<longleftrightarrow> x = -\<infinity> \<or> (\<exists>p. p < r \<and> x = extreal p)"
  10.324 -  by (cases x) auto
  10.325 -
  10.326 -lemma less_PInf_Ex_of_nat: "x \<noteq> \<infinity> \<longleftrightarrow> (\<exists>n::nat. x < extreal (real n))"
  10.327 -proof (cases x)
  10.328 -  case (real r) then show ?thesis
  10.329 -    using reals_Archimedean2[of r] by simp
  10.330 -qed simp_all
  10.331 -
  10.332 -lemma extreal_add_mono:
  10.333 -  fixes a b c d :: extreal assumes "a \<le> b" "c \<le> d" shows "a + c \<le> b + d"
  10.334 -  using assms
  10.335 -  apply (cases a)
  10.336 -  apply (cases rule: extreal3_cases[of b c d], auto)
  10.337 -  apply (cases rule: extreal3_cases[of b c d], auto)
  10.338 -  done
  10.339 -
  10.340 -lemma extreal_minus_le_minus[simp]:
  10.341 -  fixes a b :: extreal shows "- a \<le> - b \<longleftrightarrow> b \<le> a"
  10.342 -  by (cases rule: extreal2_cases[of a b]) auto
  10.343 -
  10.344 -lemma extreal_minus_less_minus[simp]:
  10.345 -  fixes a b :: extreal shows "- a < - b \<longleftrightarrow> b < a"
  10.346 -  by (cases rule: extreal2_cases[of a b]) auto
  10.347 -
  10.348 -lemma extreal_le_real_iff:
  10.349 -  "x \<le> real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x \<le> y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x \<le> 0))"
  10.350 -  by (cases y) auto
  10.351 -
  10.352 -lemma real_le_extreal_iff:
  10.353 -  "real y \<le> x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y \<le> extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 \<le> x))"
  10.354 -  by (cases y) auto
  10.355 -
  10.356 -lemma extreal_less_real_iff:
  10.357 -  "x < real y \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> extreal x < y) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> x < 0))"
  10.358 -  by (cases y) auto
  10.359 -
  10.360 -lemma real_less_extreal_iff:
  10.361 -  "real y < x \<longleftrightarrow> ((\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> y < extreal x) \<and> (\<bar>y\<bar> = \<infinity> \<longrightarrow> 0 < x))"
  10.362 -  by (cases y) auto
  10.363 -
  10.364 -lemma real_of_extreal_pos:
  10.365 -  fixes x :: extreal shows "0 \<le> x \<Longrightarrow> 0 \<le> real x" by (cases x) auto
  10.366 -
  10.367 -lemmas real_of_extreal_ord_simps =
  10.368 -  extreal_le_real_iff real_le_extreal_iff extreal_less_real_iff real_less_extreal_iff
  10.369 -
  10.370 -lemma abs_extreal_ge0[simp]: "0 \<le> x \<Longrightarrow> \<bar>x :: extreal\<bar> = x"
  10.371 -  by (cases x) auto
  10.372 -
  10.373 -lemma abs_extreal_less0[simp]: "x < 0 \<Longrightarrow> \<bar>x :: extreal\<bar> = -x"
  10.374 -  by (cases x) auto
  10.375 -
  10.376 -lemma abs_extreal_pos[simp]: "0 \<le> \<bar>x :: extreal\<bar>"
  10.377 -  by (cases x) auto
  10.378 -
  10.379 -lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \<le> 0 \<longleftrightarrow> (X \<le> 0 \<or> X = \<infinity>)"
  10.380 -  by (cases X) auto
  10.381 -
  10.382 -lemma abs_real_of_extreal[simp]: "\<bar>real (X :: extreal)\<bar> = real \<bar>X\<bar>"
  10.383 -  by (cases X) auto
  10.384 -
  10.385 -lemma zero_less_real_of_extreal: "0 < real X \<longleftrightarrow> (0 < X \<and> X \<noteq> \<infinity>)"
  10.386 -  by (cases X) auto
  10.387 -
  10.388 -lemma extreal_0_le_uminus_iff[simp]:
  10.389 -  fixes a :: extreal shows "0 \<le> -a \<longleftrightarrow> a \<le> 0"
  10.390 -  by (cases rule: extreal2_cases[of a]) auto
  10.391 -
  10.392 -lemma extreal_uminus_le_0_iff[simp]:
  10.393 -  fixes a :: extreal shows "-a \<le> 0 \<longleftrightarrow> 0 \<le> a"
  10.394 -  by (cases rule: extreal2_cases[of a]) auto
  10.395 -
  10.396 -lemma extreal_dense:
  10.397 -  fixes x y :: extreal assumes "x < y"
  10.398 -  shows "EX z. x < z & z < y"
  10.399 -proof -
  10.400 -{ assume a: "x = (-\<infinity>)"
  10.401 -  { assume "y = \<infinity>" hence ?thesis using a by (auto intro!: exI[of _ "0"]) }
  10.402 -  moreover
  10.403 -  { assume "y ~= \<infinity>"
  10.404 -    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
  10.405 -    hence ?thesis using `x < y` a by (auto intro!: exI[of _ "extreal (r - 1)"])
  10.406 -  } ultimately have ?thesis by auto
  10.407 -}
  10.408 -moreover
  10.409 -{ assume "x ~= (-\<infinity>)"
  10.410 -  with `x < y` obtain p where p: "x = extreal p" by (cases x) auto
  10.411 -  { assume "y = \<infinity>" hence ?thesis using `x < y` p
  10.412 -       by (auto intro!: exI[of _ "extreal (p + 1)"]) }
  10.413 -  moreover
  10.414 -  { assume "y ~= \<infinity>"
  10.415 -    with `x < y` obtain r where r: "y = extreal r" by (cases y) auto
  10.416 -    with p `x < y` have "p < r" by auto
  10.417 -    with dense obtain z where "p < z" "z < r" by auto
  10.418 -    hence ?thesis using r p by (auto intro!: exI[of _ "extreal z"])
  10.419 -  } ultimately have ?thesis by auto
  10.420 -} ultimately show ?thesis by auto
  10.421 -qed
  10.422 -
  10.423 -lemma extreal_dense2:
  10.424 -  fixes x y :: extreal assumes "x < y"
  10.425 -  shows "EX z. x < extreal z & extreal z < y"
  10.426 -  by (metis extreal_dense[OF `x < y`] extreal_cases less_extreal.simps(2,3))
  10.427 -
  10.428 -lemma extreal_add_strict_mono:
  10.429 -  fixes a b c d :: extreal
  10.430 -  assumes "a = b" "0 \<le> a" "a \<noteq> \<infinity>" "c < d"
  10.431 -  shows "a + c < b + d"
  10.432 -  using assms by (cases rule: extreal3_cases[case_product extreal_cases, of a b c d]) auto
  10.433 -
  10.434 -lemma extreal_less_add: "\<bar>a\<bar> \<noteq> \<infinity> \<Longrightarrow> c < b \<Longrightarrow> a + c < a + b"
  10.435 -  by (cases rule: extreal2_cases[of b c]) auto
  10.436 -
  10.437 -lemma extreal_uminus_eq_reorder: "- a = b \<longleftrightarrow> a = (-b::extreal)" by auto
  10.438 -
  10.439 -lemma extreal_uminus_less_reorder: "- a < b \<longleftrightarrow> -b < (a::extreal)"
  10.440 -  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_less_minus)
  10.441 -
  10.442 -lemma extreal_uminus_le_reorder: "- a \<le> b \<longleftrightarrow> -b \<le> (a::extreal)"
  10.443 -  by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_le_minus)
  10.444 -
  10.445 -lemmas extreal_uminus_reorder =
  10.446 -  extreal_uminus_eq_reorder extreal_uminus_less_reorder extreal_uminus_le_reorder
  10.447 -
  10.448 -lemma extreal_bot:
  10.449 -  fixes x :: extreal assumes "\<And>B. x \<le> extreal B" shows "x = - \<infinity>"
  10.450 -proof (cases x)
  10.451 -  case (real r) with assms[of "r - 1"] show ?thesis by auto
  10.452 -next case PInf with assms[of 0] show ?thesis by auto
  10.453 -next case MInf then show ?thesis by simp
  10.454 -qed
  10.455 -
  10.456 -lemma extreal_top:
  10.457 -  fixes x :: extreal assumes "\<And>B. x \<ge> extreal B" shows "x = \<infinity>"
  10.458 -proof (cases x)
  10.459 -  case (real r) with assms[of "r + 1"] show ?thesis by auto
  10.460 -next case MInf with assms[of 0] show ?thesis by auto
  10.461 -next case PInf then show ?thesis by simp
  10.462 -qed
  10.463 -
  10.464 -lemma
  10.465 -  shows extreal_max[simp]: "extreal (max x y) = max (extreal x) (extreal y)"
  10.466 -    and extreal_min[simp]: "extreal (min x y) = min (extreal x) (extreal y)"
  10.467 -  by (simp_all add: min_def max_def)
  10.468 -
  10.469 -lemma extreal_max_0: "max 0 (extreal r) = extreal (max 0 r)"
  10.470 -  by (auto simp: zero_extreal_def)
  10.471 -
  10.472 -lemma
  10.473 -  fixes f :: "nat \<Rightarrow> extreal"
  10.474 -  shows incseq_uminus[simp]: "incseq (\<lambda>x. - f x) \<longleftrightarrow> decseq f"
  10.475 -  and decseq_uminus[simp]: "decseq (\<lambda>x. - f x) \<longleftrightarrow> incseq f"
  10.476 -  unfolding decseq_def incseq_def by auto
  10.477 -
  10.478 -lemma incseq_extreal: "incseq f \<Longrightarrow> incseq (\<lambda>x. extreal (f x))"
  10.479 -  unfolding incseq_def by auto
  10.480 -
  10.481 -lemma extreal_add_nonneg_nonneg:
  10.482 -  fixes a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
  10.483 -  using add_mono[of 0 a 0 b] by simp
  10.484 -
  10.485 -lemma image_eqD: "f ` A = B \<Longrightarrow> (\<forall>x\<in>A. f x \<in> B)"
  10.486 -  by auto
  10.487 -
  10.488 -lemma incseq_setsumI:
  10.489 -  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
  10.490 -  assumes "\<And>i. 0 \<le> f i"
  10.491 -  shows "incseq (\<lambda>i. setsum f {..< i})"
  10.492 -proof (intro incseq_SucI)
  10.493 -  fix n have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
  10.494 -    using assms by (rule add_left_mono)
  10.495 -  then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
  10.496 -    by auto
  10.497 -qed
  10.498 -
  10.499 -lemma incseq_setsumI2:
  10.500 -  fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::{comm_monoid_add, ordered_ab_semigroup_add}"
  10.501 -  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
  10.502 -  shows "incseq (\<lambda>i. \<Sum>n\<in>A. f n i)"
  10.503 -  using assms unfolding incseq_def by (auto intro: setsum_mono)
  10.504 -
  10.505 -subsubsection "Multiplication"
  10.506 -
  10.507 -instantiation extreal :: "{comm_monoid_mult, sgn}"
  10.508 -begin
  10.509 -
  10.510 -definition "1 = extreal 1"
  10.511 -
  10.512 -function sgn_extreal where
  10.513 -  "sgn (extreal r) = extreal (sgn r)"
  10.514 -| "sgn \<infinity> = 1"
  10.515 -| "sgn (-\<infinity>) = -1"
  10.516 -by (auto intro: extreal_cases)
  10.517 -termination proof qed (rule wf_empty)
  10.518 -
  10.519 -function times_extreal where
  10.520 -"extreal r * extreal p = extreal (r * p)" |
  10.521 -"extreal r * \<infinity> = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
  10.522 -"\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then \<infinity> else -\<infinity>)" |
  10.523 -"extreal r * -\<infinity> = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
  10.524 -"-\<infinity> * extreal r = (if r = 0 then 0 else if r > 0 then -\<infinity> else \<infinity>)" |
  10.525 -"\<infinity> * \<infinity> = \<infinity>" |
  10.526 -"-\<infinity> * \<infinity> = -\<infinity>" |
  10.527 -"\<infinity> * -\<infinity> = -\<infinity>" |
  10.528 -"-\<infinity> * -\<infinity> = \<infinity>"
  10.529 -proof -
  10.530 -  case (goal1 P x)
  10.531 -  moreover then obtain a b where "x = (a, b)" by (cases x) auto
  10.532 -  ultimately show P by (cases rule: extreal2_cases[of a b]) auto
  10.533 -qed simp_all
  10.534 -termination by (relation "{}") simp
  10.535 -
  10.536 -instance
  10.537 -proof
  10.538 -  fix a :: extreal show "1 * a = a"
  10.539 -    by (cases a) (simp_all add: one_extreal_def)
  10.540 -  fix b :: extreal show "a * b = b * a"
  10.541 -    by (cases rule: extreal2_cases[of a b]) simp_all
  10.542 -  fix c :: extreal show "a * b * c = a * (b * c)"
  10.543 -    by (cases rule: extreal3_cases[of a b c])
  10.544 -       (simp_all add: zero_extreal_def zero_less_mult_iff)
  10.545 -qed
  10.546 -end
  10.547 -
  10.548 -lemma real_of_extreal_le_1:
  10.549 -  fixes a :: extreal shows "a \<le> 1 \<Longrightarrow> real a \<le> 1"
  10.550 -  by (cases a) (auto simp: one_extreal_def)
  10.551 -
  10.552 -lemma abs_extreal_one[simp]: "\<bar>1\<bar> = (1::extreal)"
  10.553 -  unfolding one_extreal_def by simp
  10.554 -
  10.555 -lemma extreal_mult_zero[simp]:
  10.556 -  fixes a :: extreal shows "a * 0 = 0"
  10.557 -  by (cases a) (simp_all add: zero_extreal_def)
  10.558 -
  10.559 -lemma extreal_zero_mult[simp]:
  10.560 -  fixes a :: extreal shows "0 * a = 0"
  10.561 -  by (cases a) (simp_all add: zero_extreal_def)
  10.562 -
  10.563 -lemma extreal_m1_less_0[simp]:
  10.564 -  "-(1::extreal) < 0"
  10.565 -  by (simp add: zero_extreal_def one_extreal_def)
  10.566 -
  10.567 -lemma extreal_zero_m1[simp]:
  10.568 -  "1 \<noteq> (0::extreal)"
  10.569 -  by (simp add: zero_extreal_def one_extreal_def)
  10.570 -
  10.571 -lemma extreal_times_0[simp]:
  10.572 -  fixes x :: extreal shows "0 * x = 0"
  10.573 -  by (cases x) (auto simp: zero_extreal_def)
  10.574 -
  10.575 -lemma extreal_times[simp]:
  10.576 -  "1 \<noteq> \<infinity>" "\<infinity> \<noteq> 1"
  10.577 -  "1 \<noteq> -\<infinity>" "-\<infinity> \<noteq> 1"
  10.578 -  by (auto simp add: times_extreal_def one_extreal_def)
  10.579 -
  10.580 -lemma extreal_plus_1[simp]:
  10.581 -  "1 + extreal r = extreal (r + 1)" "extreal r + 1 = extreal (r + 1)"
  10.582 -  "1 + -\<infinity> = -\<infinity>" "-\<infinity> + 1 = -\<infinity>"
  10.583 -  unfolding one_extreal_def by auto
  10.584 -
  10.585 -lemma extreal_zero_times[simp]:
  10.586 -  fixes a b :: extreal shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
  10.587 -  by (cases rule: extreal2_cases[of a b]) auto
  10.588 -
  10.589 -lemma extreal_mult_eq_PInfty[simp]:
  10.590 -  shows "a * b = \<infinity> \<longleftrightarrow>
  10.591 -    (a = \<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = -\<infinity>)"
  10.592 -  by (cases rule: extreal2_cases[of a b]) auto
  10.593 -
  10.594 -lemma extreal_mult_eq_MInfty[simp]:
  10.595 -  shows "a * b = -\<infinity> \<longleftrightarrow>
  10.596 -    (a = \<infinity> \<and> b < 0) \<or> (a < 0 \<and> b = \<infinity>) \<or> (a = -\<infinity> \<and> b > 0) \<or> (a > 0 \<and> b = -\<infinity>)"
  10.597 -  by (cases rule: extreal2_cases[of a b]) auto
  10.598 -
  10.599 -lemma extreal_0_less_1[simp]: "0 < (1::extreal)"
  10.600 -  by (simp_all add: zero_extreal_def one_extreal_def)
  10.601 -
  10.602 -lemma extreal_zero_one[simp]: "0 \<noteq> (1::extreal)"
  10.603 -  by (simp_all add: zero_extreal_def one_extreal_def)
  10.604 -
  10.605 -lemma extreal_mult_minus_left[simp]:
  10.606 -  fixes a b :: extreal shows "-a * b = - (a * b)"
  10.607 -  by (cases rule: extreal2_cases[of a b]) auto
  10.608 -
  10.609 -lemma extreal_mult_minus_right[simp]:
  10.610 -  fixes a b :: extreal shows "a * -b = - (a * b)"
  10.611 -  by (cases rule: extreal2_cases[of a b]) auto
  10.612 -
  10.613 -lemma extreal_mult_infty[simp]:
  10.614 -  "a * \<infinity> = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
  10.615 -  by (cases a) auto
  10.616 -
  10.617 -lemma extreal_infty_mult[simp]:
  10.618 -  "\<infinity> * a = (if a = 0 then 0 else if 0 < a then \<infinity> else - \<infinity>)"
  10.619 -  by (cases a) auto
  10.620 -
  10.621 -lemma extreal_mult_strict_right_mono:
  10.622 -  assumes "a < b" and "0 < c" "c < \<infinity>"
  10.623 -  shows "a * c < b * c"
  10.624 -  using assms
  10.625 -  by (cases rule: extreal3_cases[of a b c])
  10.626 -     (auto simp: zero_le_mult_iff extreal_less_PInfty)
  10.627 -
  10.628 -lemma extreal_mult_strict_left_mono:
  10.629 -  "\<lbrakk> a < b ; 0 < c ; c < \<infinity>\<rbrakk> \<Longrightarrow> c * a < c * b"
  10.630 -  using extreal_mult_strict_right_mono by (simp add: mult_commute[of c])
  10.631 -
  10.632 -lemma extreal_mult_right_mono:
  10.633 -  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> a*c \<le> b*c"
  10.634 -  using assms
  10.635 -  apply (cases "c = 0") apply simp
  10.636 -  by (cases rule: extreal3_cases[of a b c])
  10.637 -     (auto simp: zero_le_mult_iff extreal_less_PInfty)
  10.638 -
  10.639 -lemma extreal_mult_left_mono:
  10.640 -  fixes a b c :: extreal shows "\<lbrakk>a \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> c * a \<le> c * b"
  10.641 -  using extreal_mult_right_mono by (simp add: mult_commute[of c])
  10.642 -
  10.643 -lemma zero_less_one_extreal[simp]: "0 \<le> (1::extreal)"
  10.644 -  by (simp add: one_extreal_def zero_extreal_def)
  10.645 -
  10.646 -lemma extreal_0_le_mult[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * (b :: extreal)"
  10.647 -  by (cases rule: extreal2_cases[of a b]) (auto simp: mult_nonneg_nonneg)
  10.648 -
  10.649 -lemma extreal_right_distrib:
  10.650 -  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> r * (a + b) = r * a + r * b"
  10.651 -  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
  10.652 -
  10.653 -lemma extreal_left_distrib:
  10.654 -  fixes r a b :: extreal shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> (a + b) * r = a * r + b * r"
  10.655 -  by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps)
  10.656 -
  10.657 -lemma extreal_mult_le_0_iff:
  10.658 -  fixes a b :: extreal
  10.659 -  shows "a * b \<le> 0 \<longleftrightarrow> (0 \<le> a \<and> b \<le> 0) \<or> (a \<le> 0 \<and> 0 \<le> b)"
  10.660 -  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_le_0_iff)
  10.661 -
  10.662 -lemma extreal_zero_le_0_iff:
  10.663 -  fixes a b :: extreal
  10.664 -  shows "0 \<le> a * b \<longleftrightarrow> (0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0)"
  10.665 -  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_le_mult_iff)
  10.666 -
  10.667 -lemma extreal_mult_less_0_iff:
  10.668 -  fixes a b :: extreal
  10.669 -  shows "a * b < 0 \<longleftrightarrow> (0 < a \<and> b < 0) \<or> (a < 0 \<and> 0 < b)"
  10.670 -  by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_less_0_iff)
  10.671 -
  10.672 -lemma extreal_zero_less_0_iff:
  10.673 -  fixes a b :: extreal
  10.674 -  shows "0 < a * b \<longleftrightarrow> (0 < a \<and> 0 < b) \<or> (a < 0 \<and> b < 0)"
  10.675 -  by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_less_mult_iff)
  10.676 -
  10.677 -lemma extreal_distrib:
  10.678 -  fixes a b c :: extreal
  10.679 -  assumes "a \<noteq> \<infinity> \<or> b \<noteq> -\<infinity>" "a \<noteq> -\<infinity> \<or> b \<noteq> \<infinity>" "\<bar>c\<bar> \<noteq> \<infinity>"
  10.680 -  shows "(a + b) * c = a * c + b * c"
  10.681 -  using assms
  10.682 -  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
  10.683 -
  10.684 -lemma extreal_le_epsilon:
  10.685 -  fixes x y :: extreal
  10.686 -  assumes "ALL e. 0 < e --> x <= y + e"
  10.687 -  shows "x <= y"
  10.688 -proof-
  10.689 -{ assume a: "EX r. y = extreal r"
  10.690 -  from this obtain r where r_def: "y = extreal r" by auto
  10.691 -  { assume "x=(-\<infinity>)" hence ?thesis by auto }
  10.692 -  moreover
  10.693 -  { assume "~(x=(-\<infinity>))"
  10.694 -    from this obtain p where p_def: "x = extreal p"
  10.695 -    using a assms[rule_format, of 1] by (cases x) auto
  10.696 -    { fix e have "0 < e --> p <= r + e"
  10.697 -      using assms[rule_format, of "extreal e"] p_def r_def by auto }
  10.698 -    hence "p <= r" apply (subst field_le_epsilon) by auto
  10.699 -    hence ?thesis using r_def p_def by auto
  10.700 -  } ultimately have ?thesis by blast
  10.701 -}
  10.702 -moreover
  10.703 -{ assume "y=(-\<infinity>) | y=\<infinity>" hence ?thesis
  10.704 -    using assms[rule_format, of 1] by (cases x) auto
  10.705 -} ultimately show ?thesis by (cases y) auto
  10.706 -qed
  10.707 -
  10.708 -
  10.709 -lemma extreal_le_epsilon2:
  10.710 -  fixes x y :: extreal
  10.711 -  assumes "ALL e. 0 < e --> x <= y + extreal e"
  10.712 -  shows "x <= y"
  10.713 -proof-
  10.714 -{ fix e :: extreal assume "e>0"
  10.715 -  { assume "e=\<infinity>" hence "x<=y+e" by auto }
  10.716 -  moreover
  10.717 -  { assume "e~=\<infinity>"
  10.718 -    from this obtain r where "e = extreal r" using `e>0` apply (cases e) by auto
  10.719 -    hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto
  10.720 -  } ultimately have "x<=y+e" by blast
  10.721 -} from this show ?thesis using extreal_le_epsilon by auto
  10.722 -qed
  10.723 -
  10.724 -lemma extreal_le_real:
  10.725 -  fixes x y :: extreal
  10.726 -  assumes "ALL z. x <= extreal z --> y <= extreal z"
  10.727 -  shows "y <= x"
  10.728 -by (metis assms extreal.exhaust extreal_bot extreal_less_eq(1)
  10.729 -          extreal_less_eq(2) order_refl uminus_extreal.simps(2))
  10.730 -
  10.731 -lemma extreal_le_extreal:
  10.732 -  fixes x y :: extreal
  10.733 -  assumes "\<And>B. B < x \<Longrightarrow> B <= y"
  10.734 -  shows "x <= y"
  10.735 -by (metis assms extreal_dense leD linorder_le_less_linear)
  10.736 -
  10.737 -lemma extreal_ge_extreal:
  10.738 -  fixes x y :: extreal
  10.739 -  assumes "ALL B. B>x --> B >= y"
  10.740 -  shows "x >= y"
  10.741 -by (metis assms extreal_dense leD linorder_le_less_linear)
  10.742 -
  10.743 -lemma setprod_extreal_0:
  10.744 -  fixes f :: "'a \<Rightarrow> extreal"
  10.745 -  shows "(\<Prod>i\<in>A. f i) = 0 \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. f i = 0))"
  10.746 -proof cases
  10.747 -  assume "finite A"
  10.748 -  then show ?thesis by (induct A) auto
  10.749 -qed auto
  10.750 -
  10.751 -lemma setprod_extreal_pos:
  10.752 -  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)"
  10.753 -proof cases
  10.754 -  assume "finite I" from this pos show ?thesis by induct auto
  10.755 -qed simp
  10.756 -
  10.757 -lemma setprod_PInf:
  10.758 -  assumes "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
  10.759 -  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)"
  10.760 -proof cases
  10.761 -  assume "finite I" from this assms show ?thesis
  10.762 -  proof (induct I)
  10.763 -    case (insert i I)
  10.764 -    then have pos: "0 \<le> f i" "0 \<le> setprod f I" by (auto intro!: setprod_extreal_pos)
  10.765 -    from insert have "(\<Prod>j\<in>insert i I. f j) = \<infinity> \<longleftrightarrow> setprod f I * f i = \<infinity>" by auto
  10.766 -    also have "\<dots> \<longleftrightarrow> (setprod f I = \<infinity> \<or> f i = \<infinity>) \<and> f i \<noteq> 0 \<and> setprod f I \<noteq> 0"
  10.767 -      using setprod_extreal_pos[of I f] pos
  10.768 -      by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto
  10.769 -    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)"
  10.770 -      using insert by (auto simp: setprod_extreal_0)
  10.771 -    finally show ?case .
  10.772 -  qed simp
  10.773 -qed simp
  10.774 -
  10.775 -lemma setprod_extreal: "(\<Prod>i\<in>A. extreal (f i)) = extreal (setprod f A)"
  10.776 -proof cases
  10.777 -  assume "finite A" then show ?thesis
  10.778 -    by induct (auto simp: one_extreal_def)
  10.779 -qed (simp add: one_extreal_def)
  10.780 -
  10.781 -subsubsection {* Power *}
  10.782 -
  10.783 -instantiation extreal :: power
  10.784 -begin
  10.785 -primrec power_extreal where
  10.786 -  "power_extreal x 0 = 1" |
  10.787 -  "power_extreal x (Suc n) = x * x ^ n"
  10.788 -instance ..
  10.789 -end
  10.790 -
  10.791 -lemma extreal_power[simp]: "(extreal x) ^ n = extreal (x^n)"
  10.792 -  by (induct n) (auto simp: one_extreal_def)
  10.793 -
  10.794 -lemma extreal_power_PInf[simp]: "\<infinity> ^ n = (if n = 0 then 1 else \<infinity>)"
  10.795 -  by (induct n) (auto simp: one_extreal_def)
  10.796 -
  10.797 -lemma extreal_power_uminus[simp]:
  10.798 -  fixes x :: extreal
  10.799 -  shows "(- x) ^ n = (if even n then x ^ n else - (x^n))"
  10.800 -  by (induct n) (auto simp: one_extreal_def)
  10.801 -
  10.802 -lemma extreal_power_number_of[simp]:
  10.803 -  "(number_of num :: extreal) ^ n = extreal (number_of num ^ n)"
  10.804 -  by (induct n) (auto simp: one_extreal_def)
  10.805 -
  10.806 -lemma zero_le_power_extreal[simp]:
  10.807 -  fixes a :: extreal assumes "0 \<le> a"
  10.808 -  shows "0 \<le> a ^ n"
  10.809 -  using assms by (induct n) (auto simp: extreal_zero_le_0_iff)
  10.810 -
  10.811 -subsubsection {* Subtraction *}
  10.812 -
  10.813 -lemma extreal_minus_minus_image[simp]:
  10.814 -  fixes S :: "extreal set"
  10.815 -  shows "uminus ` uminus ` S = S"
  10.816 -  by (auto simp: image_iff)
  10.817 -
  10.818 -lemma extreal_uminus_lessThan[simp]:
  10.819 -  fixes a :: extreal shows "uminus ` {..<a} = {-a<..}"
  10.820 -proof (safe intro!: image_eqI)
  10.821 -  fix x assume "-a < x"
  10.822 -  then have "- x < - (- a)" by (simp del: extreal_uminus_uminus)
  10.823 -  then show "- x < a" by simp
  10.824 -qed auto
  10.825 -
  10.826 -lemma extreal_uminus_greaterThan[simp]:
  10.827 -  "uminus ` {(a::extreal)<..} = {..<-a}"
  10.828 -  by (metis extreal_uminus_lessThan extreal_uminus_uminus
  10.829 -            extreal_minus_minus_image)
  10.830 -
  10.831 -instantiation extreal :: minus
  10.832 -begin
  10.833 -definition "x - y = x + -(y::extreal)"
  10.834 -instance ..
  10.835 -end
  10.836 -
  10.837 -lemma extreal_minus[simp]:
  10.838 -  "extreal r - extreal p = extreal (r - p)"
  10.839 -  "-\<infinity> - extreal r = -\<infinity>"
  10.840 -  "extreal r - \<infinity> = -\<infinity>"
  10.841 -  "\<infinity> - x = \<infinity>"
  10.842 -  "-\<infinity> - \<infinity> = -\<infinity>"
  10.843 -  "x - -y = x + y"
  10.844 -  "x - 0 = x"
  10.845 -  "0 - x = -x"
  10.846 -  by (simp_all add: minus_extreal_def)
  10.847 -
  10.848 -lemma extreal_x_minus_x[simp]:
  10.849 -  "x - x = (if \<bar>x\<bar> = \<infinity> then \<infinity> else 0)"
  10.850 -  by (cases x) simp_all
  10.851 -
  10.852 -lemma extreal_eq_minus_iff:
  10.853 -  fixes x y z :: extreal
  10.854 -  shows "x = z - y \<longleftrightarrow>
  10.855 -    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y = z) \<and>
  10.856 -    (y = -\<infinity> \<longrightarrow> x = \<infinity>) \<and>
  10.857 -    (y = \<infinity> \<longrightarrow> z = \<infinity> \<longrightarrow> x = \<infinity>) \<and>
  10.858 -    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>)"
  10.859 -  by (cases rule: extreal3_cases[of x y z]) auto
  10.860 -
  10.861 -lemma extreal_eq_minus:
  10.862 -  fixes x y z :: extreal
  10.863 -  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x = z - y \<longleftrightarrow> x + y = z"
  10.864 -  by (auto simp: extreal_eq_minus_iff)
  10.865 -
  10.866 -lemma extreal_less_minus_iff:
  10.867 -  fixes x y z :: extreal
  10.868 -  shows "x < z - y \<longleftrightarrow>
  10.869 -    (y = \<infinity> \<longrightarrow> z = \<infinity> \<and> x \<noteq> \<infinity>) \<and>
  10.870 -    (y = -\<infinity> \<longrightarrow> x \<noteq> \<infinity>) \<and>
  10.871 -    (\<bar>y\<bar> \<noteq> \<infinity>\<longrightarrow> x + y < z)"
  10.872 -  by (cases rule: extreal3_cases[of x y z]) auto
  10.873 -
  10.874 -lemma extreal_less_minus:
  10.875 -  fixes x y z :: extreal
  10.876 -  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x < z - y \<longleftrightarrow> x + y < z"
  10.877 -  by (auto simp: extreal_less_minus_iff)
  10.878 -
  10.879 -lemma extreal_le_minus_iff:
  10.880 -  fixes x y z :: extreal
  10.881 -  shows "x \<le> z - y \<longleftrightarrow>
  10.882 -    (y = \<infinity> \<longrightarrow> z \<noteq> \<infinity> \<longrightarrow> x = -\<infinity>) \<and>
  10.883 -    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x + y \<le> z)"
  10.884 -  by (cases rule: extreal3_cases[of x y z]) auto
  10.885 -
  10.886 -lemma extreal_le_minus:
  10.887 -  fixes x y z :: extreal
  10.888 -  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x \<le> z - y \<longleftrightarrow> x + y \<le> z"
  10.889 -  by (auto simp: extreal_le_minus_iff)
  10.890 -
  10.891 -lemma extreal_minus_less_iff:
  10.892 -  fixes x y z :: extreal
  10.893 -  shows "x - y < z \<longleftrightarrow>
  10.894 -    y \<noteq> -\<infinity> \<and> (y = \<infinity> \<longrightarrow> x \<noteq> \<infinity> \<and> z \<noteq> -\<infinity>) \<and>
  10.895 -    (y \<noteq> \<infinity> \<longrightarrow> x < z + y)"
  10.896 -  by (cases rule: extreal3_cases[of x y z]) auto
  10.897 -
  10.898 -lemma extreal_minus_less:
  10.899 -  fixes x y z :: extreal
  10.900 -  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y < z \<longleftrightarrow> x < z + y"
  10.901 -  by (auto simp: extreal_minus_less_iff)
  10.902 -
  10.903 -lemma extreal_minus_le_iff:
  10.904 -  fixes x y z :: extreal
  10.905 -  shows "x - y \<le> z \<longleftrightarrow>
  10.906 -    (y = -\<infinity> \<longrightarrow> z = \<infinity>) \<and>
  10.907 -    (y = \<infinity> \<longrightarrow> x = \<infinity> \<longrightarrow> z = \<infinity>) \<and>
  10.908 -    (\<bar>y\<bar> \<noteq> \<infinity> \<longrightarrow> x \<le> z + y)"
  10.909 -  by (cases rule: extreal3_cases[of x y z]) auto
  10.910 -
  10.911 -lemma extreal_minus_le:
  10.912 -  fixes x y z :: extreal
  10.913 -  shows "\<bar>y\<bar> \<noteq> \<infinity> \<Longrightarrow> x - y \<le> z \<longleftrightarrow> x \<le> z + y"
  10.914 -  by (auto simp: extreal_minus_le_iff)
  10.915 -
  10.916 -lemma extreal_minus_eq_minus_iff:
  10.917 -  fixes a b c :: extreal
  10.918 -  shows "a - b = a - c \<longleftrightarrow>
  10.919 -    b = c \<or> a = \<infinity> \<or> (a = -\<infinity> \<and> b \<noteq> -\<infinity> \<and> c \<noteq> -\<infinity>)"
  10.920 -  by (cases rule: extreal3_cases[of a b c]) auto
  10.921 -
  10.922 -lemma extreal_add_le_add_iff:
  10.923 -  "c + a \<le> c + b \<longleftrightarrow>
  10.924 -    a \<le> b \<or> c = \<infinity> \<or> (c = -\<infinity> \<and> a \<noteq> \<infinity> \<and> b \<noteq> \<infinity>)"
  10.925 -  by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps)
  10.926 -
  10.927 -lemma extreal_mult_le_mult_iff:
  10.928 -  "\<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)"
  10.929 -  by (cases rule: extreal3_cases[of a b c]) (simp_all add: mult_le_cancel_left)
  10.930 -
  10.931 -lemma extreal_minus_mono:
  10.932 -  fixes A B C D :: extreal assumes "A \<le> B" "D \<le> C"
  10.933 -  shows "A - C \<le> B - D"
  10.934 -  using assms
  10.935 -  by (cases rule: extreal3_cases[case_product extreal_cases, of A B C D]) simp_all
  10.936 -
  10.937 -lemma real_of_extreal_minus:
  10.938 -  "real (a - b) = (if \<bar>a\<bar> = \<infinity> \<or> \<bar>b\<bar> = \<infinity> then 0 else real a - real b)"
  10.939 -  by (cases rule: extreal2_cases[of a b]) auto
  10.940 -
  10.941 -lemma extreal_diff_positive:
  10.942 -  fixes a b :: extreal shows "a \<le> b \<Longrightarrow> 0 \<le> b - a"
  10.943 -  by (cases rule: extreal2_cases[of a b]) auto
  10.944 -
  10.945 -lemma extreal_between:
  10.946 -  fixes x e :: extreal
  10.947 -  assumes "\<bar>x\<bar> \<noteq> \<infinity>" "0 < e"
  10.948 -  shows "x - e < x" "x < x + e"
  10.949 -using assms apply (cases x, cases e) apply auto
  10.950 -using assms by (cases x, cases e) auto
  10.951 -
  10.952 -subsubsection {* Division *}
  10.953 -
  10.954 -instantiation extreal :: inverse
  10.955 -begin
  10.956 -
  10.957 -function inverse_extreal where
  10.958 -"inverse (extreal r) = (if r = 0 then \<infinity> else extreal (inverse r))" |
  10.959 -"inverse \<infinity> = 0" |
  10.960 -"inverse (-\<infinity>) = 0"
  10.961 -  by (auto intro: extreal_cases)
  10.962 -termination by (relation "{}") simp
  10.963 -
  10.964 -definition "x / y = x * inverse (y :: extreal)"
  10.965 -
  10.966 -instance proof qed
  10.967 -end
  10.968 -
  10.969 -lemma real_of_extreal_inverse[simp]:
  10.970 -  fixes a :: extreal
  10.971 -  shows "real (inverse a) = 1 / real a"
  10.972 -  by (cases a) (auto simp: inverse_eq_divide)
  10.973 -
  10.974 -lemma extreal_inverse[simp]:
  10.975 -  "inverse 0 = \<infinity>"
  10.976 -  "inverse (1::extreal) = 1"
  10.977 -  by (simp_all add: one_extreal_def zero_extreal_def)
  10.978 -
  10.979 -lemma extreal_divide[simp]:
  10.980 -  "extreal r / extreal p = (if p = 0 then extreal r * \<infinity> else extreal (r / p))"
  10.981 -  unfolding divide_extreal_def by (auto simp: divide_real_def)
  10.982 -
  10.983 -lemma extreal_divide_same[simp]:
  10.984 -  "x / x = (if \<bar>x\<bar> = \<infinity> \<or> x = 0 then 0 else 1)"
  10.985 -  by (cases x)
  10.986 -     (simp_all add: divide_real_def divide_extreal_def one_extreal_def)
  10.987 -
  10.988 -lemma extreal_inv_inv[simp]:
  10.989 -  "inverse (inverse x) = (if x \<noteq> -\<infinity> then x else \<infinity>)"
  10.990 -  by (cases x) auto
  10.991 -
  10.992 -lemma extreal_inverse_minus[simp]:
  10.993 -  "inverse (- x) = (if x = 0 then \<infinity> else -inverse x)"
  10.994 -  by (cases x) simp_all
  10.995 -
  10.996 -lemma extreal_uminus_divide[simp]:
  10.997 -  fixes x y :: extreal shows "- x / y = - (x / y)"
  10.998 -  unfolding divide_extreal_def by simp
  10.999 -
 10.1000 -lemma extreal_divide_Infty[simp]:
 10.1001 -  "x / \<infinity> = 0" "x / -\<infinity> = 0"
 10.1002 -  unfolding divide_extreal_def by simp_all
 10.1003 -
 10.1004 -lemma extreal_divide_one[simp]:
 10.1005 -  "x / 1 = (x::extreal)"
 10.1006 -  unfolding divide_extreal_def by simp
 10.1007 -
 10.1008 -lemma extreal_divide_extreal[simp]:
 10.1009 -  "\<infinity> / extreal r = (if 0 \<le> r then \<infinity> else -\<infinity>)"
 10.1010 -  unfolding divide_extreal_def by simp
 10.1011 -
 10.1012 -lemma zero_le_divide_extreal[simp]:
 10.1013 -  fixes a :: extreal assumes "0 \<le> a" "0 \<le> b"
 10.1014 -  shows "0 \<le> a / b"
 10.1015 -  using assms by (cases rule: extreal2_cases[of a b]) (auto simp: zero_le_divide_iff)
 10.1016 -
 10.1017 -lemma extreal_le_divide_pos:
 10.1018 -  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> x * y \<le> z"
 10.1019 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.1020 -
 10.1021 -lemma extreal_divide_le_pos:
 10.1022 -  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> z \<le> x * y"
 10.1023 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.1024 -
 10.1025 -lemma extreal_le_divide_neg:
 10.1026 -  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> y \<le> z / x \<longleftrightarrow> z \<le> x * y"
 10.1027 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.1028 -
 10.1029 -lemma extreal_divide_le_neg:
 10.1030 -  "x < 0 \<Longrightarrow> x \<noteq> -\<infinity> \<Longrightarrow> z / x \<le> y \<longleftrightarrow> x * y \<le> z"
 10.1031 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.1032 -
 10.1033 -lemma extreal_inverse_antimono_strict:
 10.1034 -  fixes x y :: extreal
 10.1035 -  shows "0 \<le> x \<Longrightarrow> x < y \<Longrightarrow> inverse y < inverse x"
 10.1036 -  by (cases rule: extreal2_cases[of x y]) auto
 10.1037 -
 10.1038 -lemma extreal_inverse_antimono:
 10.1039 -  fixes x y :: extreal
 10.1040 -  shows "0 \<le> x \<Longrightarrow> x <= y \<Longrightarrow> inverse y <= inverse x"
 10.1041 -  by (cases rule: extreal2_cases[of x y]) auto
 10.1042 -
 10.1043 -lemma inverse_inverse_Pinfty_iff[simp]:
 10.1044 -  "inverse x = \<infinity> \<longleftrightarrow> x = 0"
 10.1045 -  by (cases x) auto
 10.1046 -
 10.1047 -lemma extreal_inverse_eq_0:
 10.1048 -  "inverse x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity>"
 10.1049 -  by (cases x) auto
 10.1050 -
 10.1051 -lemma extreal_0_gt_inverse:
 10.1052 -  fixes x :: extreal shows "0 < inverse x \<longleftrightarrow> x \<noteq> \<infinity> \<and> 0 \<le> x"
 10.1053 -  by (cases x) auto
 10.1054 -
 10.1055 -lemma extreal_mult_less_right:
 10.1056 -  assumes "b * a < c * a" "0 < a" "a < \<infinity>"
 10.1057 -  shows "b < c"
 10.1058 -  using assms
 10.1059 -  by (cases rule: extreal3_cases[of a b c])
 10.1060 -     (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
 10.1061 -
 10.1062 -lemma extreal_power_divide:
 10.1063 -  "y \<noteq> 0 \<Longrightarrow> (x / y :: extreal) ^ n = x^n / y^n"
 10.1064 -  by (cases rule: extreal2_cases[of x y])
 10.1065 -     (auto simp: one_extreal_def zero_extreal_def power_divide not_le
 10.1066 -                 power_less_zero_eq zero_le_power_iff)
 10.1067 -
 10.1068 -lemma extreal_le_mult_one_interval:
 10.1069 -  fixes x y :: extreal
 10.1070 -  assumes y: "y \<noteq> -\<infinity>"
 10.1071 -  assumes z: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
 10.1072 -  shows "x \<le> y"
 10.1073 -proof (cases x)
 10.1074 -  case PInf with z[of "1 / 2"] show "x \<le> y" by (simp add: one_extreal_def)
 10.1075 -next
 10.1076 -  case (real r) note r = this
 10.1077 -  show "x \<le> y"
 10.1078 -  proof (cases y)
 10.1079 -    case (real p) note p = this
 10.1080 -    have "r \<le> p"
 10.1081 -    proof (rule field_le_mult_one_interval)
 10.1082 -      fix z :: real assume "0 < z" and "z < 1"
 10.1083 -      with z[of "extreal z"]
 10.1084 -      show "z * r \<le> p" using p r by (auto simp: zero_le_mult_iff one_extreal_def)
 10.1085 -    qed
 10.1086 -    then show "x \<le> y" using p r by simp
 10.1087 -  qed (insert y, simp_all)
 10.1088 -qed simp
 10.1089 -
 10.1090 -subsection "Complete lattice"
 10.1091 -
 10.1092 -instantiation extreal :: lattice
 10.1093 -begin
 10.1094 -definition [simp]: "sup x y = (max x y :: extreal)"
 10.1095 -definition [simp]: "inf x y = (min x y :: extreal)"
 10.1096 -instance proof qed simp_all
 10.1097 -end
 10.1098 -
 10.1099 -instantiation extreal :: complete_lattice
 10.1100 -begin
 10.1101 -
 10.1102 -definition "bot = -\<infinity>"
 10.1103 -definition "top = \<infinity>"
 10.1104 -
 10.1105 -definition "Sup S = (LEAST z. ALL x:S. x <= z :: extreal)"
 10.1106 -definition "Inf S = (GREATEST z. ALL x:S. z <= x :: extreal)"
 10.1107 -
 10.1108 -lemma extreal_complete_Sup:
 10.1109 -  fixes S :: "extreal set" assumes "S \<noteq> {}"
 10.1110 -  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
 10.1111 -proof cases
 10.1112 -  assume "\<exists>x. \<forall>a\<in>S. a \<le> extreal x"
 10.1113 -  then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> extreal y" by auto
 10.1114 -  then have "\<infinity> \<notin> S" by force
 10.1115 -  show ?thesis
 10.1116 -  proof cases
 10.1117 -    assume "S = {-\<infinity>}"
 10.1118 -    then show ?thesis by (auto intro!: exI[of _ "-\<infinity>"])
 10.1119 -  next
 10.1120 -    assume "S \<noteq> {-\<infinity>}"
 10.1121 -    with `S \<noteq> {}` `\<infinity> \<notin> S` obtain x where "x \<in> S - {-\<infinity>}" "x \<noteq> \<infinity>" by auto
 10.1122 -    with y `\<infinity> \<notin> S` have "\<forall>z\<in>real ` (S - {-\<infinity>}). z \<le> y"
 10.1123 -      by (auto simp: real_of_extreal_ord_simps)
 10.1124 -    with reals_complete2[of "real ` (S - {-\<infinity>})"] `x \<in> S - {-\<infinity>}`
 10.1125 -    obtain s where s:
 10.1126 -       "\<forall>y\<in>S - {-\<infinity>}. real y \<le> s" "\<And>z. (\<forall>y\<in>S - {-\<infinity>}. real y \<le> z) \<Longrightarrow> s \<le> z"
 10.1127 -       by auto
 10.1128 -    show ?thesis
 10.1129 -    proof (safe intro!: exI[of _ "extreal s"])
 10.1130 -      fix z assume "z \<in> S" with `\<infinity> \<notin> S` show "z \<le> extreal s"
 10.1131 -      proof (cases z)
 10.1132 -        case (real r)
 10.1133 -        then show ?thesis
 10.1134 -          using s(1)[rule_format, of z] `z \<in> S` `z = extreal r` by auto
 10.1135 -      qed auto
 10.1136 -    next
 10.1137 -      fix z assume *: "\<forall>y\<in>S. y \<le> z"
 10.1138 -      with `S \<noteq> {-\<infinity>}` `S \<noteq> {}` show "extreal s \<le> z"
 10.1139 -      proof (cases z)
 10.1140 -        case (real u)
 10.1141 -        with * have "s \<le> u"
 10.1142 -          by (intro s(2)[of u]) (auto simp: real_of_extreal_ord_simps)
 10.1143 -        then show ?thesis using real by simp
 10.1144 -      qed auto
 10.1145 -    qed
 10.1146 -  qed
 10.1147 -next
 10.1148 -  assume *: "\<not> (\<exists>x. \<forall>a\<in>S. a \<le> extreal x)"
 10.1149 -  show ?thesis
 10.1150 -  proof (safe intro!: exI[of _ \<infinity>])
 10.1151 -    fix y assume **: "\<forall>z\<in>S. z \<le> y"
 10.1152 -    with * show "\<infinity> \<le> y"
 10.1153 -    proof (cases y)
 10.1154 -      case MInf with * ** show ?thesis by (force simp: not_le)
 10.1155 -    qed auto
 10.1156 -  qed simp
 10.1157 -qed
 10.1158 -
 10.1159 -lemma extreal_complete_Inf:
 10.1160 -  fixes S :: "extreal set" assumes "S ~= {}"
 10.1161 -  shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)"
 10.1162 -proof-
 10.1163 -def S1 == "uminus ` S"
 10.1164 -hence "S1 ~= {}" using assms by auto
 10.1165 -from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)"
 10.1166 -   using extreal_complete_Sup[of S1] by auto
 10.1167 -{ fix z assume "ALL y:S. z <= y"
 10.1168 -  hence "ALL y:S1. y <= -z" unfolding S1_def by auto
 10.1169 -  hence "x <= -z" using x_def by auto
 10.1170 -  hence "z <= -x"
 10.1171 -    apply (subst extreal_uminus_uminus[symmetric])
 10.1172 -    unfolding extreal_minus_le_minus . }
 10.1173 -moreover have "(ALL y:S. -x <= y)"
 10.1174 -   using x_def unfolding S1_def
 10.1175 -   apply simp
 10.1176 -   apply (subst (3) extreal_uminus_uminus[symmetric])
 10.1177 -   unfolding extreal_minus_le_minus by simp
 10.1178 -ultimately show ?thesis by auto
 10.1179 -qed
 10.1180 -
 10.1181 -lemma extreal_complete_uminus_eq:
 10.1182 -  fixes S :: "extreal set"
 10.1183 -  shows "(\<forall>y\<in>uminus`S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>uminus`S. y \<le> z) \<longrightarrow> x \<le> z)
 10.1184 -     \<longleftrightarrow> (\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> -x)"
 10.1185 -  by simp (metis extreal_minus_le_minus extreal_uminus_uminus)
 10.1186 -
 10.1187 -lemma extreal_Sup_uminus_image_eq:
 10.1188 -  fixes S :: "extreal set"
 10.1189 -  shows "Sup (uminus ` S) = - Inf S"
 10.1190 -proof cases
 10.1191 -  assume "S = {}"
 10.1192 -  moreover have "(THE x. All (op \<le> x)) = (-\<infinity>::extreal)"
 10.1193 -    by (rule the_equality) (auto intro!: extreal_bot)
 10.1194 -  moreover have "(SOME x. \<forall>y. y \<le> x) = (\<infinity>::extreal)"
 10.1195 -    by (rule some_equality) (auto intro!: extreal_top)
 10.1196 -  ultimately show ?thesis unfolding Inf_extreal_def Sup_extreal_def
 10.1197 -    Least_def Greatest_def GreatestM_def by simp
 10.1198 -next
 10.1199 -  assume "S \<noteq> {}"
 10.1200 -  with extreal_complete_Sup[of "uminus`S"]
 10.1201 -  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)"
 10.1202 -    unfolding extreal_complete_uminus_eq by auto
 10.1203 -  show "Sup (uminus ` S) = - Inf S"
 10.1204 -    unfolding Inf_extreal_def Greatest_def GreatestM_def
 10.1205 -  proof (intro someI2[of _ _ "\<lambda>x. Sup (uminus`S) = - x"])
 10.1206 -    show "(\<forall>y\<in>S. -x \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> -x)"
 10.1207 -      using x .
 10.1208 -    fix x' assume "(\<forall>y\<in>S. x' \<le> y) \<and> (\<forall>y. (\<forall>z\<in>S. y \<le> z) \<longrightarrow> y \<le> x')"
 10.1209 -    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)"
 10.1210 -      unfolding extreal_complete_uminus_eq by simp
 10.1211 -    then show "Sup (uminus ` S) = -x'"
 10.1212 -      unfolding Sup_extreal_def extreal_uminus_eq_iff
 10.1213 -      by (intro Least_equality) auto
 10.1214 -  qed
 10.1215 -qed
 10.1216 -
 10.1217 -instance
 10.1218 -proof
 10.1219 -  { fix x :: extreal and A
 10.1220 -    show "bot <= x" by (cases x) (simp_all add: bot_extreal_def)
 10.1221 -    show "x <= top" by (simp add: top_extreal_def) }
 10.1222 -
 10.1223 -  { fix x :: extreal and A assume "x : A"
 10.1224 -    with extreal_complete_Sup[of A]
 10.1225 -    obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
 10.1226 -    hence "x <= s" using `x : A` by auto
 10.1227 -    also have "... = Sup A" using s unfolding Sup_extreal_def
 10.1228 -      by (auto intro!: Least_equality[symmetric])
 10.1229 -    finally show "x <= Sup A" . }
 10.1230 -  note le_Sup = this
 10.1231 -
 10.1232 -  { fix x :: extreal and A assume *: "!!z. (z : A ==> z <= x)"
 10.1233 -    show "Sup A <= x"
 10.1234 -    proof (cases "A = {}")
 10.1235 -      case True
 10.1236 -      hence "Sup A = -\<infinity>" unfolding Sup_extreal_def
 10.1237 -        by (auto intro!: Least_equality)
 10.1238 -      thus "Sup A <= x" by simp
 10.1239 -    next
 10.1240 -      case False
 10.1241 -      with extreal_complete_Sup[of A]
 10.1242 -      obtain s where s: "\<forall>y\<in>A. y <= s" "\<forall>z. (\<forall>y\<in>A. y <= z) \<longrightarrow> s <= z" by auto
 10.1243 -      hence "Sup A = s"
 10.1244 -        unfolding Sup_extreal_def by (auto intro!: Least_equality)
 10.1245 -      also have "s <= x" using * s by auto
 10.1246 -      finally show "Sup A <= x" .
 10.1247 -    qed }
 10.1248 -  note Sup_le = this
 10.1249 -
 10.1250 -  { fix x :: extreal and A assume "x \<in> A"
 10.1251 -    with le_Sup[of "-x" "uminus`A"] show "Inf A \<le> x"
 10.1252 -      unfolding extreal_Sup_uminus_image_eq by simp }
 10.1253 -
 10.1254 -  { fix x :: extreal and A assume *: "!!z. (z : A ==> x <= z)"
 10.1255 -    with Sup_le[of "uminus`A" "-x"] show "x \<le> Inf A"
 10.1256 -      unfolding extreal_Sup_uminus_image_eq by force }
 10.1257 -qed
 10.1258 -end
 10.1259 -
 10.1260 -lemma extreal_SUPR_uminus:
 10.1261 -  fixes f :: "'a => extreal"
 10.1262 -  shows "(SUP i : R. -(f i)) = -(INF i : R. f i)"
 10.1263 -  unfolding SUPR_def INFI_def
 10.1264 -  using extreal_Sup_uminus_image_eq[of "f`R"]
 10.1265 -  by (simp add: image_image)
 10.1266 -
 10.1267 -lemma extreal_INFI_uminus:
 10.1268 -  fixes f :: "'a => extreal"
 10.1269 -  shows "(INF i : R. -(f i)) = -(SUP i : R. f i)"
 10.1270 -  using extreal_SUPR_uminus[of _ "\<lambda>x. - f x"] by simp
 10.1271 -
 10.1272 -lemma extreal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::extreal set)"
 10.1273 -  using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image)
 10.1274 -
 10.1275 -lemma extreal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: extreal set)"
 10.1276 -  by (auto intro!: inj_onI)
 10.1277 -
 10.1278 -lemma extreal_image_uminus_shift:
 10.1279 -  fixes X Y :: "extreal set" shows "uminus ` X = Y \<longleftrightarrow> X = uminus ` Y"
 10.1280 -proof
 10.1281 -  assume "uminus ` X = Y"
 10.1282 -  then have "uminus ` uminus ` X = uminus ` Y"
 10.1283 -    by (simp add: inj_image_eq_iff)
 10.1284 -  then show "X = uminus ` Y" by (simp add: image_image)
 10.1285 -qed (simp add: image_image)
 10.1286 -
 10.1287 -lemma Inf_extreal_iff:
 10.1288 -  fixes z :: extreal
 10.1289 -  shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x<y) <-> Inf X < y"
 10.1290 -  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
 10.1291 -            order_less_le_trans)
 10.1292 -
 10.1293 -lemma Sup_eq_MInfty:
 10.1294 -  fixes S :: "extreal set" shows "Sup S = -\<infinity> \<longleftrightarrow> S = {} \<or> S = {-\<infinity>}"
 10.1295 -proof
 10.1296 -  assume a: "Sup S = -\<infinity>"
 10.1297 -  with complete_lattice_class.Sup_upper[of _ S]
 10.1298 -  show "S={} \<or> S={-\<infinity>}" by auto
 10.1299 -next
 10.1300 -  assume "S={} \<or> S={-\<infinity>}" then show "Sup S = -\<infinity>"
 10.1301 -    unfolding Sup_extreal_def by (auto intro!: Least_equality)
 10.1302 -qed
 10.1303 -
 10.1304 -lemma Inf_eq_PInfty:
 10.1305 -  fixes S :: "extreal set" shows "Inf S = \<infinity> \<longleftrightarrow> S = {} \<or> S = {\<infinity>}"
 10.1306 -  using Sup_eq_MInfty[of "uminus`S"]
 10.1307 -  unfolding extreal_Sup_uminus_image_eq extreal_image_uminus_shift by simp
 10.1308 -
 10.1309 -lemma Inf_eq_MInfty: "-\<infinity> : S ==> Inf S = -\<infinity>"
 10.1310 -  unfolding Inf_extreal_def
 10.1311 -  by (auto intro!: Greatest_equality)
 10.1312 -
 10.1313 -lemma Sup_eq_PInfty: "\<infinity> : S ==> Sup S = \<infinity>"
 10.1314 -  unfolding Sup_extreal_def
 10.1315 -  by (auto intro!: Least_equality)
 10.1316 -
 10.1317 -lemma extreal_SUPI:
 10.1318 -  fixes x :: extreal
 10.1319 -  assumes "!!i. i : A ==> f i <= x"
 10.1320 -  assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y"
 10.1321 -  shows "(SUP i:A. f i) = x"
 10.1322 -  unfolding SUPR_def Sup_extreal_def
 10.1323 -  using assms by (auto intro!: Least_equality)
 10.1324 -
 10.1325 -lemma extreal_INFI:
 10.1326 -  fixes x :: extreal
 10.1327 -  assumes "!!i. i : A ==> f i >= x"
 10.1328 -  assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y"
 10.1329 -  shows "(INF i:A. f i) = x"
 10.1330 -  unfolding INFI_def Inf_extreal_def
 10.1331 -  using assms by (auto intro!: Greatest_equality)
 10.1332 -
 10.1333 -lemma Sup_extreal_close:
 10.1334 -  fixes e :: extreal
 10.1335 -  assumes "0 < e" and S: "\<bar>Sup S\<bar> \<noteq> \<infinity>" "S \<noteq> {}"
 10.1336 -  shows "\<exists>x\<in>S. Sup S - e < x"
 10.1337 -  using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1])
 10.1338 -
 10.1339 -lemma Inf_extreal_close:
 10.1340 -  fixes e :: extreal assumes "\<bar>Inf X\<bar> \<noteq> \<infinity>" "0 < e"
 10.1341 -  shows "\<exists>x\<in>X. x < Inf X + e"
 10.1342 -proof (rule Inf_less_iff[THEN iffD1])
 10.1343 -  show "Inf X < Inf X + e" using assms
 10.1344 -    by (cases e) auto
 10.1345 -qed
 10.1346 -
 10.1347 -lemma Sup_eq_top_iff:
 10.1348 -  fixes A :: "'a::{complete_lattice, linorder} set"
 10.1349 -  shows "Sup A = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < i)"
 10.1350 -proof
 10.1351 -  assume *: "Sup A = top"
 10.1352 -  show "(\<forall>x<top. \<exists>i\<in>A. x < i)" unfolding *[symmetric]
 10.1353 -  proof (intro allI impI)
 10.1354 -    fix x assume "x < Sup A" then show "\<exists>i\<in>A. x < i"
 10.1355 -      unfolding less_Sup_iff by auto
 10.1356 -  qed
 10.1357 -next
 10.1358 -  assume *: "\<forall>x<top. \<exists>i\<in>A. x < i"
 10.1359 -  show "Sup A = top"
 10.1360 -  proof (rule ccontr)
 10.1361 -    assume "Sup A \<noteq> top"
 10.1362 -    with top_greatest[of "Sup A"]
 10.1363 -    have "Sup A < top" unfolding le_less by auto
 10.1364 -    then have "Sup A < Sup A"
 10.1365 -      using * unfolding less_Sup_iff by auto
 10.1366 -    then show False by auto
 10.1367 -  qed
 10.1368 -qed
 10.1369 -
 10.1370 -lemma SUP_eq_top_iff:
 10.1371 -  fixes f :: "'a \<Rightarrow> 'b::{complete_lattice, linorder}"
 10.1372 -  shows "(SUP i:A. f i) = top \<longleftrightarrow> (\<forall>x<top. \<exists>i\<in>A. x < f i)"
 10.1373 -  unfolding SUPR_def Sup_eq_top_iff by auto
 10.1374 -
 10.1375 -lemma SUP_nat_Infty: "(SUP i::nat. extreal (real i)) = \<infinity>"
 10.1376 -proof -
 10.1377 -  { fix x assume "x \<noteq> \<infinity>"
 10.1378 -    then have "\<exists>k::nat. x < extreal (real k)"
 10.1379 -    proof (cases x)
 10.1380 -      case MInf then show ?thesis by (intro exI[of _ 0]) auto
 10.1381 -    next
 10.1382 -      case (real r)
 10.1383 -      moreover obtain k :: nat where "r < real k"
 10.1384 -        using ex_less_of_nat by (auto simp: real_eq_of_nat)
 10.1385 -      ultimately show ?thesis by auto
 10.1386 -    qed simp }
 10.1387 -  then show ?thesis
 10.1388 -    using SUP_eq_top_iff[of UNIV "\<lambda>n::nat. extreal (real n)"]
 10.1389 -    by (auto simp: top_extreal_def)
 10.1390 -qed
 10.1391 -
 10.1392 -lemma extreal_le_Sup:
 10.1393 -  fixes x :: extreal
 10.1394 -  shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))"
 10.1395 -(is "?lhs <-> ?rhs")
 10.1396 -proof-
 10.1397 -{ assume "?rhs"
 10.1398 -  { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i)<x" by (simp add: not_le)
 10.1399 -    from this obtain y where y_def: "(SUP i:A. f i)<y & y<x" using extreal_dense by auto
 10.1400 -    from this obtain i where "i : A & y <= f i" using `?rhs` by auto
 10.1401 -    hence "y <= (SUP i:A. f i)" using le_SUPI[of i A f] by auto
 10.1402 -    hence False using y_def by auto
 10.1403 -  } hence "?lhs" by auto
 10.1404 -}
 10.1405 -moreover
 10.1406 -{ assume "?lhs" hence "?rhs"
 10.1407 -  by (metis Collect_def Collect_mem_eq SUP_leI assms atLeastatMost_empty atLeastatMost_empty_iff
 10.1408 -      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
 10.1409 -} ultimately show ?thesis by auto
 10.1410 -qed
 10.1411 -
 10.1412 -lemma extreal_Inf_le:
 10.1413 -  fixes x :: extreal
 10.1414 -  shows "((INF i:A. f i) <= x) <-> (ALL y. x < y --> (EX i. i : A & f i <= y))"
 10.1415 -(is "?lhs <-> ?rhs")
 10.1416 -proof-
 10.1417 -{ assume "?rhs"
 10.1418 -  { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le)
 10.1419 -    from this obtain y where y_def: "x<y & y<(INF i:A. f i)" using extreal_dense by auto
 10.1420 -    from this obtain i where "i : A & f i <= y" using `?rhs` by auto
 10.1421 -    hence "(INF i:A. f i) <= y" using INF_leI[of i A f] by auto
 10.1422 -    hence False using y_def by auto
 10.1423 -  } hence "?lhs" by auto
 10.1424 -}
 10.1425 -moreover
 10.1426 -{ assume "?lhs" hence "?rhs"
 10.1427 -  by (metis Collect_def Collect_mem_eq le_INFI assms atLeastatMost_empty atLeastatMost_empty_iff
 10.1428 -      inf_sup_ord(4) linorder_le_cases sup_absorb1 xt1(8))
 10.1429 -} ultimately show ?thesis by auto
 10.1430 -qed
 10.1431 -
 10.1432 -lemma Inf_less:
 10.1433 -  fixes x :: extreal
 10.1434 -  assumes "(INF i:A. f i) < x"
 10.1435 -  shows "EX i. i : A & f i <= x"
 10.1436 -proof(rule ccontr)
 10.1437 -  assume "~ (EX i. i : A & f i <= x)"
 10.1438 -  hence "ALL i:A. f i > x" by auto
 10.1439 -  hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto
 10.1440 -  thus False using assms by auto
 10.1441 -qed
 10.1442 -
 10.1443 -lemma same_INF:
 10.1444 -  assumes "ALL e:A. f e = g e"
 10.1445 -  shows "(INF e:A. f e) = (INF e:A. g e)"
 10.1446 -proof-
 10.1447 -have "f ` A = g ` A" unfolding image_def using assms by auto
 10.1448 -thus ?thesis unfolding INFI_def by auto
 10.1449 -qed
 10.1450 -
 10.1451 -lemma same_SUP:
 10.1452 -  assumes "ALL e:A. f e = g e"
 10.1453 -  shows "(SUP e:A. f e) = (SUP e:A. g e)"
 10.1454 -proof-
 10.1455 -have "f ` A = g ` A" unfolding image_def using assms by auto
 10.1456 -thus ?thesis unfolding SUPR_def by auto
 10.1457 -qed
 10.1458 -
 10.1459 -lemma SUPR_eq:
 10.1460 -  assumes "\<forall>i\<in>A. \<exists>j\<in>B. f i \<le> g j"
 10.1461 -  assumes "\<forall>j\<in>B. \<exists>i\<in>A. g j \<le> f i"
 10.1462 -  shows "(SUP i:A. f i) = (SUP j:B. g j)"
 10.1463 -proof (intro antisym)
 10.1464 -  show "(SUP i:A. f i) \<le> (SUP j:B. g j)"
 10.1465 -    using assms by (metis SUP_leI le_SUPI2)
 10.1466 -  show "(SUP i:B. g i) \<le> (SUP j:A. f j)"
 10.1467 -    using assms by (metis SUP_leI le_SUPI2)
 10.1468 -qed
 10.1469 -
 10.1470 -lemma SUP_extreal_le_addI:
 10.1471 -  assumes "\<And>i. f i + y \<le> z" and "y \<noteq> -\<infinity>"
 10.1472 -  shows "SUPR UNIV f + y \<le> z"
 10.1473 -proof (cases y)
 10.1474 -  case (real r)
 10.1475 -  then have "\<And>i. f i \<le> z - y" using assms by (simp add: extreal_le_minus_iff)
 10.1476 -  then have "SUPR UNIV f \<le> z - y" by (rule SUP_leI)
 10.1477 -  then show ?thesis using real by (simp add: extreal_le_minus_iff)
 10.1478 -qed (insert assms, auto)
 10.1479 -
 10.1480 -lemma SUPR_extreal_add:
 10.1481 -  fixes f g :: "nat \<Rightarrow> extreal"
 10.1482 -  assumes "incseq f" "incseq g" and pos: "\<And>i. f i \<noteq> -\<infinity>" "\<And>i. g i \<noteq> -\<infinity>"
 10.1483 -  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
 10.1484 -proof (rule extreal_SUPI)
 10.1485 -  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> f i + g i \<le> y"
 10.1486 -  have f: "SUPR UNIV f \<noteq> -\<infinity>" using pos
 10.1487 -    unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD)
 10.1488 -  { fix j
 10.1489 -    { fix i
 10.1490 -      have "f i + g j \<le> f i + g (max i j)"
 10.1491 -        using `incseq g`[THEN incseqD] by (rule add_left_mono) auto
 10.1492 -      also have "\<dots> \<le> f (max i j) + g (max i j)"
 10.1493 -        using `incseq f`[THEN incseqD] by (rule add_right_mono) auto
 10.1494 -      also have "\<dots> \<le> y" using * by auto
 10.1495 -      finally have "f i + g j \<le> y" . }
 10.1496 -    then have "SUPR UNIV f + g j \<le> y"
 10.1497 -      using assms(4)[of j] by (intro SUP_extreal_le_addI) auto
 10.1498 -    then have "g j + SUPR UNIV f \<le> y" by (simp add: ac_simps) }
 10.1499 -  then have "SUPR UNIV g + SUPR UNIV f \<le> y"
 10.1500 -    using f by (rule SUP_extreal_le_addI)
 10.1501 -  then show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
 10.1502 -qed (auto intro!: add_mono le_SUPI)
 10.1503 -
 10.1504 -lemma SUPR_extreal_add_pos:
 10.1505 -  fixes f g :: "nat \<Rightarrow> extreal"
 10.1506 -  assumes inc: "incseq f" "incseq g" and pos: "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
 10.1507 -  shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g"
 10.1508 -proof (intro SUPR_extreal_add inc)
 10.1509 -  fix i show "f i \<noteq> -\<infinity>" "g i \<noteq> -\<infinity>" using pos[of i] by auto
 10.1510 -qed
 10.1511 -
 10.1512 -lemma SUPR_extreal_setsum:
 10.1513 -  fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> extreal"
 10.1514 -  assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)" and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
 10.1515 -  shows "(SUP i. \<Sum>n\<in>A. f n i) = (\<Sum>n\<in>A. SUPR UNIV (f n))"
 10.1516 -proof cases
 10.1517 -  assume "finite A" then show ?thesis using assms
 10.1518 -    by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_extreal_add_pos)
 10.1519 -qed simp
 10.1520 -
 10.1521 -lemma SUPR_extreal_cmult:
 10.1522 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" "0 \<le> c"
 10.1523 -  shows "(SUP i. c * f i) = c * SUPR UNIV f"
 10.1524 -proof (rule extreal_SUPI)
 10.1525 -  fix i have "f i \<le> SUPR UNIV f" by (rule le_SUPI) auto
 10.1526 -  then show "c * f i \<le> c * SUPR UNIV f"
 10.1527 -    using `0 \<le> c` by (rule extreal_mult_left_mono)
 10.1528 -next
 10.1529 -  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c * f i \<le> y"
 10.1530 -  show "c * SUPR UNIV f \<le> y"
 10.1531 -  proof cases
 10.1532 -    assume c: "0 < c \<and> c \<noteq> \<infinity>"
 10.1533 -    with * have "SUPR UNIV f \<le> y / c"
 10.1534 -      by (intro SUP_leI) (auto simp: extreal_le_divide_pos)
 10.1535 -    with c show ?thesis
 10.1536 -      by (auto simp: extreal_le_divide_pos)
 10.1537 -  next
 10.1538 -    { assume "c = \<infinity>" have ?thesis
 10.1539 -      proof cases
 10.1540 -        assume "\<forall>i. f i = 0"
 10.1541 -        moreover then have "range f = {0}" by auto
 10.1542 -        ultimately show "c * SUPR UNIV f \<le> y" using * by (auto simp: SUPR_def)
 10.1543 -      next
 10.1544 -        assume "\<not> (\<forall>i. f i = 0)"
 10.1545 -        then obtain i where "f i \<noteq> 0" by auto
 10.1546 -        with *[of i] `c = \<infinity>` `0 \<le> f i` show ?thesis by (auto split: split_if_asm)
 10.1547 -      qed }
 10.1548 -    moreover assume "\<not> (0 < c \<and> c \<noteq> \<infinity>)"
 10.1549 -    ultimately show ?thesis using * `0 \<le> c` by auto
 10.1550 -  qed
 10.1551 -qed
 10.1552 -
 10.1553 -lemma SUP_PInfty:
 10.1554 -  fixes f :: "'a \<Rightarrow> extreal"
 10.1555 -  assumes "\<And>n::nat. \<exists>i\<in>A. extreal (real n) \<le> f i"
 10.1556 -  shows "(SUP i:A. f i) = \<infinity>"
 10.1557 -  unfolding SUPR_def Sup_eq_top_iff[where 'a=extreal, unfolded top_extreal_def]
 10.1558 -  apply simp
 10.1559 -proof safe
 10.1560 -  fix x assume "x \<noteq> \<infinity>"
 10.1561 -  show "\<exists>i\<in>A. x < f i"
 10.1562 -  proof (cases x)
 10.1563 -    case PInf with `x \<noteq> \<infinity>` show ?thesis by simp
 10.1564 -  next
 10.1565 -    case MInf with assms[of "0"] show ?thesis by force
 10.1566 -  next
 10.1567 -    case (real r)
 10.1568 -    with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < extreal (real n)" by auto
 10.1569 -    moreover from assms[of n] guess i ..
 10.1570 -    ultimately show ?thesis
 10.1571 -      by (auto intro!: bexI[of _ i])
 10.1572 -  qed
 10.1573 -qed
 10.1574 -
 10.1575 -lemma Sup_countable_SUPR:
 10.1576 -  assumes "A \<noteq> {}"
 10.1577 -  shows "\<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
 10.1578 -proof (cases "Sup A")
 10.1579 -  case (real r)
 10.1580 -  have "\<forall>n::nat. \<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
 10.1581 -  proof
 10.1582 -    fix n ::nat have "\<exists>x\<in>A. Sup A - 1 / extreal (real n) < x"
 10.1583 -      using assms real by (intro Sup_extreal_close) (auto simp: one_extreal_def)
 10.1584 -    then guess x ..
 10.1585 -    then show "\<exists>x. x \<in> A \<and> Sup A < x + 1 / extreal (real n)"
 10.1586 -      by (auto intro!: exI[of _ x] simp: extreal_minus_less_iff)
 10.1587 -  qed
 10.1588 -  from choice[OF this] guess f .. note f = this
 10.1589 -  have "SUPR UNIV f = Sup A"
 10.1590 -  proof (rule extreal_SUPI)
 10.1591 -    fix i show "f i \<le> Sup A" using f
 10.1592 -      by (auto intro!: complete_lattice_class.Sup_upper)
 10.1593 -  next
 10.1594 -    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
 10.1595 -    show "Sup A \<le> y"
 10.1596 -    proof (rule extreal_le_epsilon, intro allI impI)
 10.1597 -      fix e :: extreal assume "0 < e"
 10.1598 -      show "Sup A \<le> y + e"
 10.1599 -      proof (cases e)
 10.1600 -        case (real r)
 10.1601 -        hence "0 < r" using `0 < e` by auto
 10.1602 -        then obtain n ::nat where *: "1 / real n < r" "0 < n"
 10.1603 -          using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide)
 10.1604 -        have "Sup A \<le> f n + 1 / extreal (real n)" using f[THEN spec, of n] by auto
 10.1605 -        also have "1 / extreal (real n) \<le> e" using real * by (auto simp: one_extreal_def )
 10.1606 -        with bound have "f n + 1 / extreal (real n) \<le> y + e" by (rule add_mono) simp
 10.1607 -        finally show "Sup A \<le> y + e" .
 10.1608 -      qed (insert `0 < e`, auto)
 10.1609 -    qed
 10.1610 -  qed
 10.1611 -  with f show ?thesis by (auto intro!: exI[of _ f])
 10.1612 -next
 10.1613 -  case PInf
 10.1614 -  from `A \<noteq> {}` obtain x where "x \<in> A" by auto
 10.1615 -  show ?thesis
 10.1616 -  proof cases
 10.1617 -    assume "\<infinity> \<in> A"
 10.1618 -    moreover then have "\<infinity> \<le> Sup A" by (intro complete_lattice_class.Sup_upper)
 10.1619 -    ultimately show ?thesis by (auto intro!: exI[of _ "\<lambda>x. \<infinity>"])
 10.1620 -  next
 10.1621 -    assume "\<infinity> \<notin> A"
 10.1622 -    have "\<exists>x\<in>A. 0 \<le> x"
 10.1623 -      by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least extreal_infty_less_eq2 linorder_linear)
 10.1624 -    then obtain x where "x \<in> A" "0 \<le> x" by auto
 10.1625 -    have "\<forall>n::nat. \<exists>f. f \<in> A \<and> x + extreal (real n) \<le> f"
 10.1626 -    proof (rule ccontr)
 10.1627 -      assume "\<not> ?thesis"
 10.1628 -      then have "\<exists>n::nat. Sup A \<le> x + extreal (real n)"
 10.1629 -        by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le)
 10.1630 -      then show False using `x \<in> A` `\<infinity> \<notin> A` PInf
 10.1631 -        by(cases x) auto
 10.1632 -    qed
 10.1633 -    from choice[OF this] guess f .. note f = this
 10.1634 -    have "SUPR UNIV f = \<infinity>"
 10.1635 -    proof (rule SUP_PInfty)
 10.1636 -      fix n :: nat show "\<exists>i\<in>UNIV. extreal (real n) \<le> f i"
 10.1637 -        using f[THEN spec, of n] `0 \<le> x`
 10.1638 -        by (cases rule: extreal2_cases[of "f n" x]) (auto intro!: exI[of _ n])
 10.1639 -    qed
 10.1640 -    then show ?thesis using f PInf by (auto intro!: exI[of _ f])
 10.1641 -  qed
 10.1642 -next
 10.1643 -  case MInf
 10.1644 -  with `A \<noteq> {}` have "A = {-\<infinity>}" by (auto simp: Sup_eq_MInfty)
 10.1645 -  then show ?thesis using MInf by (auto intro!: exI[of _ "\<lambda>x. -\<infinity>"])
 10.1646 -qed
 10.1647 -
 10.1648 -lemma SUPR_countable_SUPR:
 10.1649 -  "A \<noteq> {} \<Longrightarrow> \<exists>f::nat \<Rightarrow> extreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
 10.1650 -  using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def)
 10.1651 -
 10.1652 -
 10.1653 -lemma Sup_extreal_cadd:
 10.1654 -  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
 10.1655 -  shows "Sup ((\<lambda>x. a + x) ` A) = a + Sup A"
 10.1656 -proof (rule antisym)
 10.1657 -  have *: "\<And>a::extreal. \<And>A. Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A"
 10.1658 -    by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
 10.1659 -  then show "Sup ((\<lambda>x. a + x) ` A) \<le> a + Sup A" .
 10.1660 -  show "a + Sup A \<le> Sup ((\<lambda>x. a + x) ` A)"
 10.1661 -  proof (cases a)
 10.1662 -    case PInf with `A \<noteq> {}` show ?thesis by (auto simp: image_constant)
 10.1663 -  next
 10.1664 -    case (real r)
 10.1665 -    then have **: "op + (- a) ` op + a ` A = A"
 10.1666 -      by (auto simp: image_iff ac_simps zero_extreal_def[symmetric])
 10.1667 -    from *[of "-a" "(\<lambda>x. a + x) ` A"] real show ?thesis unfolding **
 10.1668 -      by (cases rule: extreal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto
 10.1669 -  qed (insert `a \<noteq> -\<infinity>`, auto)
 10.1670 -qed
 10.1671 -
 10.1672 -lemma Sup_extreal_cminus:
 10.1673 -  fixes A :: "extreal set" assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
 10.1674 -  shows "Sup ((\<lambda>x. a - x) ` A) = a - Inf A"
 10.1675 -  using Sup_extreal_cadd[of "uminus ` A" a] assms
 10.1676 -  by (simp add: comp_def image_image minus_extreal_def
 10.1677 -                 extreal_Sup_uminus_image_eq)
 10.1678 -
 10.1679 -lemma SUPR_extreal_cminus:
 10.1680 -  fixes A assumes "A \<noteq> {}" and "a \<noteq> -\<infinity>"
 10.1681 -  shows "(SUP x:A. a - f x) = a - (INF x:A. f x)"
 10.1682 -  using Sup_extreal_cminus[of "f`A" a] assms
 10.1683 -  unfolding SUPR_def INFI_def image_image by auto
 10.1684 -
 10.1685 -lemma Inf_extreal_cminus:
 10.1686 -  fixes A :: "extreal set" assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
 10.1687 -  shows "Inf ((\<lambda>x. a - x) ` A) = a - Sup A"
 10.1688 -proof -
 10.1689 -  { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto }
 10.1690 -  moreover then have "(\<lambda>x. -a - x)`uminus`A = uminus ` (\<lambda>x. a - x) ` A"
 10.1691 -    by (auto simp: image_image)
 10.1692 -  ultimately show ?thesis
 10.1693 -    using Sup_extreal_cminus[of "uminus ` A" "-a"] assms
 10.1694 -    by (auto simp add: extreal_Sup_uminus_image_eq extreal_Inf_uminus_image_eq)
 10.1695 -qed
 10.1696 -
 10.1697 -lemma INFI_extreal_cminus:
 10.1698 -  fixes A assumes "A \<noteq> {}" and "\<bar>a\<bar> \<noteq> \<infinity>"
 10.1699 -  shows "(INF x:A. a - f x) = a - (SUP x:A. f x)"
 10.1700 -  using Inf_extreal_cminus[of "f`A" a] assms
 10.1701 -  unfolding SUPR_def INFI_def image_image
 10.1702 -  by auto
 10.1703 -
 10.1704 -lemma uminus_extreal_add_uminus_uminus:
 10.1705 -  fixes a b :: extreal shows "a \<noteq> \<infinity> \<Longrightarrow> b \<noteq> \<infinity> \<Longrightarrow> - (- a + - b) = a + b"
 10.1706 -  by (cases rule: extreal2_cases[of a b]) auto
 10.1707 -
 10.1708 -lemma INFI_extreal_add:
 10.1709 -  assumes "decseq f" "decseq g" and fin: "\<And>i. f i \<noteq> \<infinity>" "\<And>i. g i \<noteq> \<infinity>"
 10.1710 -  shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g"
 10.1711 -proof -
 10.1712 -  have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
 10.1713 -    using assms unfolding INF_less_iff by auto
 10.1714 -  { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i"
 10.1715 -      by (rule uminus_extreal_add_uminus_uminus) }
 10.1716 -  then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
 10.1717 -    by simp
 10.1718 -  also have "\<dots> = INFI UNIV f + INFI UNIV g"
 10.1719 -    unfolding extreal_INFI_uminus
 10.1720 -    using assms INF_less
 10.1721 -    by (subst SUPR_extreal_add)
 10.1722 -       (auto simp: extreal_SUPR_uminus intro!: uminus_extreal_add_uminus_uminus)
 10.1723 -  finally show ?thesis .
 10.1724 -qed
 10.1725 -
 10.1726 -subsection "Limits on @{typ extreal}"
 10.1727 -
 10.1728 -subsubsection "Topological space"
 10.1729 -
 10.1730 -instantiation extreal :: topological_space
 10.1731 -begin
 10.1732 -
 10.1733 -definition "open A \<longleftrightarrow> open (extreal -` A)
 10.1734 -       \<and> (\<infinity> \<in> A \<longrightarrow> (\<exists>x. {extreal x <..} \<subseteq> A))
 10.1735 -       \<and> (-\<infinity> \<in> A \<longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A))"
 10.1736 -
 10.1737 -lemma open_PInfty: "open A \<Longrightarrow> \<infinity> \<in> A \<Longrightarrow> (\<exists>x. {extreal x<..} \<subseteq> A)"
 10.1738 -  unfolding open_extreal_def by auto
 10.1739 -
 10.1740 -lemma open_MInfty: "open A \<Longrightarrow> -\<infinity> \<in> A \<Longrightarrow> (\<exists>x. {..<extreal x} \<subseteq> A)"
 10.1741 -  unfolding open_extreal_def by auto
 10.1742 -
 10.1743 -lemma open_PInfty2: assumes "open A" "\<infinity> \<in> A" obtains x where "{extreal x<..} \<subseteq> A"
 10.1744 -  using open_PInfty[OF assms] by auto
 10.1745 -
 10.1746 -lemma open_MInfty2: assumes "open A" "-\<infinity> \<in> A" obtains x where "{..<extreal x} \<subseteq> A"
 10.1747 -  using open_MInfty[OF assms] by auto
 10.1748 -
 10.1749 -lemma extreal_openE: assumes "open A" obtains x y where
 10.1750 -  "open (extreal -` A)"
 10.1751 -  "\<infinity> \<in> A \<Longrightarrow> {extreal x<..} \<subseteq> A"
 10.1752 -  "-\<infinity> \<in> A \<Longrightarrow> {..<extreal y} \<subseteq> A"
 10.1753 -  using assms open_extreal_def by auto
 10.1754 -
 10.1755 -instance
 10.1756 -proof
 10.1757 -  let ?U = "UNIV::extreal set"
 10.1758 -  show "open ?U" unfolding open_extreal_def
 10.1759 -    by (auto intro!: exI[of _ 0])
 10.1760 -next
 10.1761 -  fix S T::"extreal set" assume "open S" and "open T"
 10.1762 -  from `open S`[THEN extreal_openE] guess xS yS .
 10.1763 -  moreover from `open T`[THEN extreal_openE] guess xT yT .
 10.1764 -  ultimately have
 10.1765 -    "open (extreal -` (S \<inter> T))"
 10.1766 -    "\<infinity> \<in> S \<inter> T \<Longrightarrow> {extreal (max xS xT) <..} \<subseteq> S \<inter> T"
 10.1767 -    "-\<infinity> \<in> S \<inter> T \<Longrightarrow> {..< extreal (min yS yT)} \<subseteq> S \<inter> T"
 10.1768 -    by auto
 10.1769 -  then show "open (S Int T)" unfolding open_extreal_def by blast
 10.1770 -next
 10.1771 -  fix K :: "extreal set set" assume "\<forall>S\<in>K. open S"
 10.1772 -  then have *: "\<forall>S. \<exists>x y. S \<in> K \<longrightarrow> open (extreal -` S) \<and>
 10.1773 -    (\<infinity> \<in> S \<longrightarrow> {extreal x <..} \<subseteq> S) \<and> (-\<infinity> \<in> S \<longrightarrow> {..< extreal y} \<subseteq> S)"
 10.1774 -    by (auto simp: open_extreal_def)
 10.1775 -  then show "open (Union K)" unfolding open_extreal_def
 10.1776 -  proof (intro conjI impI)
 10.1777 -    show "open (extreal -` \<Union>K)"
 10.1778 -      using *[THEN choice] by (auto simp: vimage_Union)
 10.1779 -  qed ((metis UnionE Union_upper subset_trans *)+)
 10.1780 -qed
 10.1781 -end
 10.1782 -
 10.1783 -lemma open_extreal: "open S \<Longrightarrow> open (extreal ` S)"
 10.1784 -  by (auto simp: inj_vimage_image_eq open_extreal_def)
 10.1785 -
 10.1786 -lemma open_extreal_vimage: "open S \<Longrightarrow> open (extreal -` S)"
 10.1787 -  unfolding open_extreal_def by auto
 10.1788 -
 10.1789 -lemma open_extreal_lessThan[intro, simp]: "open {..< a :: extreal}"
 10.1790 -proof -
 10.1791 -  have "\<And>x. extreal -` {..<extreal x} = {..< x}"
 10.1792 -    "extreal -` {..< \<infinity>} = UNIV" "extreal -` {..< -\<infinity>} = {}" by auto
 10.1793 -  then show ?thesis by (cases a) (auto simp: open_extreal_def)
 10.1794 -qed
 10.1795 -
 10.1796 -lemma open_extreal_greaterThan[intro, simp]:
 10.1797 -  "open {a :: extreal <..}"
 10.1798 -proof -
 10.1799 -  have "\<And>x. extreal -` {extreal x<..} = {x<..}"
 10.1800 -    "extreal -` {\<infinity><..} = {}" "extreal -` {-\<infinity><..} = UNIV" by auto
 10.1801 -  then show ?thesis by (cases a) (auto simp: open_extreal_def)
 10.1802 -qed
 10.1803 -
 10.1804 -lemma extreal_open_greaterThanLessThan[intro, simp]: "open {a::extreal <..< b}"
 10.1805 -  unfolding greaterThanLessThan_def by auto
 10.1806 -
 10.1807 -lemma closed_extreal_atLeast[simp, intro]: "closed {a :: extreal ..}"
 10.1808 -proof -
 10.1809 -  have "- {a ..} = {..< a}" by auto
 10.1810 -  then show "closed {a ..}"
 10.1811 -    unfolding closed_def using open_extreal_lessThan by auto
 10.1812 -qed
 10.1813 -
 10.1814 -lemma closed_extreal_atMost[simp, intro]: "closed {.. b :: extreal}"
 10.1815 -proof -
 10.1816 -  have "- {.. b} = {b <..}" by auto
 10.1817 -  then show "closed {.. b}"
 10.1818 -    unfolding closed_def using open_extreal_greaterThan by auto
 10.1819 -qed
 10.1820 -
 10.1821 -lemma closed_extreal_atLeastAtMost[simp, intro]:
 10.1822 -  shows "closed {a :: extreal .. b}"
 10.1823 -  unfolding atLeastAtMost_def by auto
 10.1824 -
 10.1825 -lemma closed_extreal_singleton:
 10.1826 -  "closed {a :: extreal}"
 10.1827 -by (metis atLeastAtMost_singleton closed_extreal_atLeastAtMost)
 10.1828 -
 10.1829 -lemma extreal_open_cont_interval:
 10.1830 -  assumes "open S" "x \<in> S" "\<bar>x\<bar> \<noteq> \<infinity>"
 10.1831 -  obtains e where "e>0" "{x-e <..< x+e} \<subseteq> S"
 10.1832 -proof-
 10.1833 -  from `open S` have "open (extreal -` S)" by (rule extreal_openE)
 10.1834 -  then obtain e where "0 < e" and e: "\<And>y. dist y (real x) < e \<Longrightarrow> extreal y \<in> S"
 10.1835 -    using assms unfolding open_dist by force
 10.1836 -  show thesis
 10.1837 -  proof (intro that subsetI)
 10.1838 -    show "0 < extreal e" using `0 < e` by auto
 10.1839 -    fix y assume "y \<in> {x - extreal e<..<x + extreal e}"
 10.1840 -    with assms obtain t where "y = extreal t" "dist t (real x) < e"
 10.1841 -      apply (cases y) by (auto simp: dist_real_def)
 10.1842 -    then show "y \<in> S" using e[of t] by auto
 10.1843 -  qed
 10.1844 -qed
 10.1845 -
 10.1846 -lemma extreal_open_cont_interval2:
 10.1847 -  assumes "open S" "x \<in> S" and x: "\<bar>x\<bar> \<noteq> \<infinity>"
 10.1848 -  obtains a b where "a < x" "x < b" "{a <..< b} \<subseteq> S"
 10.1849 -proof-
 10.1850 -  guess e using extreal_open_cont_interval[OF assms] .
 10.1851 -  with that[of "x-e" "x+e"] extreal_between[OF x, of e]
 10.1852 -  show thesis by auto
 10.1853 -qed
 10.1854 -
 10.1855 -instance extreal :: t2_space
 10.1856 -proof
 10.1857 -  fix x y :: extreal assume "x ~= y"
 10.1858 -  let "?P x (y::extreal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}"
 10.1859 -
 10.1860 -  { fix x y :: extreal assume "x < y"
 10.1861 -    from extreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
 10.1862 -    have "?P x y"
 10.1863 -      apply (rule exI[of _ "{..<z}"])
 10.1864 -      apply (rule exI[of _ "{z<..}"])
 10.1865 -      using z by auto }
 10.1866 -  note * = this
 10.1867 -
 10.1868 -  from `x ~= y`
 10.1869 -  show "EX U V. open U & open V & x : U & y : V & U Int V = {}"
 10.1870 -  proof (cases rule: linorder_cases)
 10.1871 -    assume "x = y" with `x ~= y` show ?thesis by simp
 10.1872 -  next assume "x < y" from *[OF this] show ?thesis by auto
 10.1873 -  next assume "y < x" from *[OF this] show ?thesis by auto
 10.1874 -  qed
 10.1875 -qed
 10.1876 -
 10.1877 -subsubsection {* Convergent sequences *}
 10.1878 -
 10.1879 -lemma lim_extreal[simp]:
 10.1880 -  "((\<lambda>n. extreal (f n)) ---> extreal x) net \<longleftrightarrow> (f ---> x) net" (is "?l = ?r")
 10.1881 -proof (intro iffI topological_tendstoI)
 10.1882 -  fix S assume "?l" "open S" "x \<in> S"
 10.1883 -  then show "eventually (\<lambda>x. f x \<in> S) net"
 10.1884 -    using `?l`[THEN topological_tendstoD, OF open_extreal, OF `open S`]
 10.1885 -    by (simp add: inj_image_mem_iff)
 10.1886 -next
 10.1887 -  fix S assume "?r" "open S" "extreal x \<in> S"
 10.1888 -  show "eventually (\<lambda>x. extreal (f x) \<in> S) net"
 10.1889 -    using `?r`[THEN topological_tendstoD, OF open_extreal_vimage, OF `open S`]
 10.1890 -    using `extreal x \<in> S` by auto
 10.1891 -qed
 10.1892 -
 10.1893 -lemma lim_real_of_extreal[simp]:
 10.1894 -  assumes lim: "(f ---> extreal x) net"
 10.1895 -  shows "((\<lambda>x. real (f x)) ---> x) net"
 10.1896 -proof (intro topological_tendstoI)
 10.1897 -  fix S assume "open S" "x \<in> S"
 10.1898 -  then have S: "open S" "extreal x \<in> extreal ` S"
 10.1899 -    by (simp_all add: inj_image_mem_iff)
 10.1900 -  have "\<forall>x. f x \<in> extreal ` S \<longrightarrow> real (f x) \<in> S" by auto
 10.1901 -  from this lim[THEN topological_tendstoD, OF open_extreal, OF S]
 10.1902 -  show "eventually (\<lambda>x. real (f x) \<in> S) net"
 10.1903 -    by (rule eventually_mono)
 10.1904 -qed
 10.1905 -
 10.1906 -lemma Lim_PInfty: "f ----> \<infinity> <-> (ALL B. EX N. ALL n>=N. f n >= extreal B)" (is "?l = ?r")
 10.1907 -proof assume ?r show ?l apply(rule topological_tendstoI)
 10.1908 -    unfolding eventually_sequentially
 10.1909 -  proof- fix S assume "open S" "\<infinity> : S"
 10.1910 -    from open_PInfty[OF this] guess B .. note B=this
 10.1911 -    from `?r`[rule_format,of "B+1"] guess N .. note N=this
 10.1912 -    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
 10.1913 -    proof safe case goal1
 10.1914 -      have "extreal B < extreal (B + 1)" by auto
 10.1915 -      also have "... <= f n" using goal1 N by auto
 10.1916 -      finally show ?case using B by fastsimp
 10.1917 -    qed
 10.1918 -  qed
 10.1919 -next assume ?l show ?r
 10.1920 -  proof fix B::real have "open {extreal B<..}" "\<infinity> : {extreal B<..}" by auto
 10.1921 -    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
 10.1922 -    guess N .. note N=this
 10.1923 -    show "EX N. ALL n>=N. extreal B <= f n" apply(rule_tac x=N in exI) using N by auto
 10.1924 -  qed
 10.1925 -qed
 10.1926 -
 10.1927 -
 10.1928 -lemma Lim_MInfty: "f ----> (-\<infinity>) <-> (ALL B. EX N. ALL n>=N. f n <= extreal B)" (is "?l = ?r")
 10.1929 -proof assume ?r show ?l apply(rule topological_tendstoI)
 10.1930 -    unfolding eventually_sequentially
 10.1931 -  proof- fix S assume "open S" "(-\<infinity>) : S"
 10.1932 -    from open_MInfty[OF this] guess B .. note B=this
 10.1933 -    from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this
 10.1934 -    show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI)
 10.1935 -    proof safe case goal1
 10.1936 -      have "extreal (B - 1) >= f n" using goal1 N by auto
 10.1937 -      also have "... < extreal B" by auto
 10.1938 -      finally show ?case using B by fastsimp
 10.1939 -    qed
 10.1940 -  qed
 10.1941 -next assume ?l show ?r
 10.1942 -  proof fix B::real have "open {..<extreal B}" "(-\<infinity>) : {..<extreal B}" by auto
 10.1943 -    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
 10.1944 -    guess N .. note N=this
 10.1945 -    show "EX N. ALL n>=N. extreal B >= f n" apply(rule_tac x=N in exI) using N by auto
 10.1946 -  qed
 10.1947 -qed
 10.1948 -
 10.1949 -
 10.1950 -lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= extreal B" shows "l ~= \<infinity>"
 10.1951 -proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\<infinity>"
 10.1952 -  from lim[unfolded this Lim_PInfty,rule_format,of "?B"]
 10.1953 -  guess N .. note N=this[rule_format,OF le_refl]
 10.1954 -  hence "extreal ?B <= extreal B" using assms(2)[of N] by(rule order_trans)
 10.1955 -  hence "extreal ?B < extreal ?B" apply (rule le_less_trans) by auto
 10.1956 -  thus False by auto
 10.1957 -qed
 10.1958 -
 10.1959 -
 10.1960 -lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= extreal B" shows "l ~= (-\<infinity>)"
 10.1961 -proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\<infinity>)"
 10.1962 -  from lim[unfolded this Lim_MInfty,rule_format,of "?B"]
 10.1963 -  guess N .. note N=this[rule_format,OF le_refl]
 10.1964 -  hence "extreal B <= extreal ?B" using assms(2)[of N] order_trans[of "extreal B" "f N" "extreal(B - 1)"] by blast
 10.1965 -  thus False by auto
 10.1966 -qed
 10.1967 -
 10.1968 -
 10.1969 -lemma tendsto_explicit:
 10.1970 -  "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))"
 10.1971 -  unfolding tendsto_def eventually_sequentially by auto
 10.1972 -
 10.1973 -
 10.1974 -lemma tendsto_obtains_N:
 10.1975 -  assumes "f ----> f0"
 10.1976 -  assumes "open S" "f0 : S"
 10.1977 -  obtains N where "ALL n>=N. f n : S"
 10.1978 -  using tendsto_explicit[of f f0] assms by auto
 10.1979 -
 10.1980 -
 10.1981 -lemma tail_same_limit:
 10.1982 -  fixes X Y N
 10.1983 -  assumes "X ----> L" "ALL n>=N. X n = Y n"
 10.1984 -  shows "Y ----> L"
 10.1985 -proof-
 10.1986 -{ fix S assume "open S" and "L:S"
 10.1987 -  from this obtain N1 where "ALL n>=N1. X n : S"
 10.1988 -     using assms unfolding tendsto_def eventually_sequentially by auto
 10.1989 -  hence "ALL n>=max N N1. Y n : S" using assms by auto
 10.1990 -  hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto
 10.1991 -}
 10.1992 -thus ?thesis using tendsto_explicit by auto
 10.1993 -qed
 10.1994 -
 10.1995 -
 10.1996 -lemma Lim_bounded_PInfty2:
 10.1997 -assumes lim:"f ----> l" and "ALL n>=N. f n <= extreal B"
 10.1998 -shows "l ~= \<infinity>"
 10.1999 -proof-
 10.2000 -  def g == "(%n. if n>=N then f n else extreal B)"
 10.2001 -  hence "g ----> l" using tail_same_limit[of f l N g] lim by auto
 10.2002 -  moreover have "!!n. g n <= extreal B" using g_def assms by auto
 10.2003 -  ultimately show ?thesis using  Lim_bounded_PInfty by auto
 10.2004 -qed
 10.2005 -
 10.2006 -lemma Lim_bounded_extreal:
 10.2007 -  assumes lim:"f ----> (l :: extreal)"
 10.2008 -  and "ALL n>=M. f n <= C"
 10.2009 -  shows "l<=C"
 10.2010 -proof-
 10.2011 -{ assume "l=(-\<infinity>)" hence ?thesis by auto }
 10.2012 -moreover
 10.2013 -{ assume "~(l=(-\<infinity>))"
 10.2014 -  { assume "C=\<infinity>" hence ?thesis by auto }
 10.2015 -  moreover
 10.2016 -  { assume "C=(-\<infinity>)" hence "ALL n>=M. f n = (-\<infinity>)" using assms by auto
 10.2017 -    hence "l=(-\<infinity>)" using assms
 10.2018 -       tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\<lambda>n. -\<infinity>" "-\<infinity>" M f, OF tendsto_const] by auto
 10.2019 -    hence ?thesis by auto }
 10.2020 -  moreover
 10.2021 -  { assume "EX B. C = extreal B"
 10.2022 -    from this obtain B where B_def: "C=extreal B" by auto
 10.2023 -    hence "~(l=\<infinity>)" using Lim_bounded_PInfty2 assms by auto
 10.2024 -    from this obtain m where m_def: "extreal m=l" using `~(l=(-\<infinity>))` by (cases l) auto
 10.2025 -    from this obtain N where N_def: "ALL n>=N. f n : {extreal(m - 1) <..< extreal(m+1)}"
 10.2026 -       apply (subst tendsto_obtains_N[of f l "{extreal(m - 1) <..< extreal(m+1)}"]) using assms by auto
 10.2027 -    { fix n assume "n>=N"
 10.2028 -      hence "EX r. extreal r = f n" using N_def by (cases "f n") auto
 10.2029 -    } from this obtain g where g_def: "ALL n>=N. extreal (g n) = f n" by metis
 10.2030 -    hence "(%n. extreal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto
 10.2031 -    hence *: "(%n. g n) ----> m" using m_def by auto
 10.2032 -    { fix n assume "n>=max N M"
 10.2033 -      hence "extreal (g n) <= extreal B" using assms g_def B_def by auto
 10.2034 -      hence "g n <= B" by auto
 10.2035 -    } hence "EX N. ALL n>=N. g n <= B" by blast
 10.2036 -    hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto
 10.2037 -    hence ?thesis using m_def B_def by auto
 10.2038 -  } ultimately have ?thesis by (cases C) auto
 10.2039 -} ultimately show ?thesis by blast
 10.2040 -qed
 10.2041 -
 10.2042 -lemma real_of_extreal_mult[simp]:
 10.2043 -  fixes a b :: extreal shows "real (a * b) = real a * real b"
 10.2044 -  by (cases rule: extreal2_cases[of a b]) auto
 10.2045 -
 10.2046 -lemma real_of_extreal_eq_0:
 10.2047 -  "real x = 0 \<longleftrightarrow> x = \<infinity> \<or> x = -\<infinity> \<or> x = 0"
 10.2048 -  by (cases x) auto
 10.2049 -
 10.2050 -lemma tendsto_extreal_realD:
 10.2051 -  fixes f :: "'a \<Rightarrow> extreal"
 10.2052 -  assumes "x \<noteq> 0" and tendsto: "((\<lambda>x. extreal (real (f x))) ---> x) net"
 10.2053 -  shows "(f ---> x) net"
 10.2054 -proof (intro topological_tendstoI)
 10.2055 -  fix S assume S: "open S" "x \<in> S"
 10.2056 -  with `x \<noteq> 0` have "open (S - {0})" "x \<in> S - {0}" by auto
 10.2057 -  from tendsto[THEN topological_tendstoD, OF this]
 10.2058 -  show "eventually (\<lambda>x. f x \<in> S) net"
 10.2059 -    by (rule eventually_rev_mp) (auto simp: extreal_real real_of_extreal_0)
 10.2060 -qed
 10.2061 -
 10.2062 -lemma tendsto_extreal_realI:
 10.2063 -  fixes f :: "'a \<Rightarrow> extreal"
 10.2064 -  assumes x: "\<bar>x\<bar> \<noteq> \<infinity>" and tendsto: "(f ---> x) net"
 10.2065 -  shows "((\<lambda>x. extreal (real (f x))) ---> x) net"
 10.2066 -proof (intro topological_tendstoI)
 10.2067 -  fix S assume "open S" "x \<in> S"
 10.2068 -  with x have "open (S - {\<infinity>, -\<infinity>})" "x \<in> S - {\<infinity>, -\<infinity>}" by auto
 10.2069 -  from tendsto[THEN topological_tendstoD, OF this]
 10.2070 -  show "eventually (\<lambda>x. extreal (real (f x)) \<in> S) net"
 10.2071 -    by (elim eventually_elim1) (auto simp: extreal_real)
 10.2072 -qed
 10.2073 -
 10.2074 -lemma extreal_mult_cancel_left:
 10.2075 -  fixes a b c :: extreal shows "a * b = a * c \<longleftrightarrow>
 10.2076 -    ((\<bar>a\<bar> = \<infinity> \<and> 0 < b * c) \<or> a = 0 \<or> b = c)"
 10.2077 -  by (cases rule: extreal3_cases[of a b c])
 10.2078 -     (simp_all add: zero_less_mult_iff)
 10.2079 -
 10.2080 -lemma extreal_inj_affinity:
 10.2081 -  assumes "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" "\<bar>t\<bar> \<noteq> \<infinity>"
 10.2082 -  shows "inj_on (\<lambda>x. m * x + t) A"
 10.2083 -  using assms
 10.2084 -  by (cases rule: extreal2_cases[of m t])
 10.2085 -     (auto intro!: inj_onI simp: extreal_add_cancel_right extreal_mult_cancel_left)
 10.2086 -
 10.2087 -lemma extreal_PInfty_eq_plus[simp]:
 10.2088 -  shows "\<infinity> = a + b \<longleftrightarrow> a = \<infinity> \<or> b = \<infinity>"
 10.2089 -  by (cases rule: extreal2_cases[of a b]) auto
 10.2090 -
 10.2091 -lemma extreal_MInfty_eq_plus[simp]:
 10.2092 -  shows "-\<infinity> = a + b \<longleftrightarrow> (a = -\<infinity> \<and> b \<noteq> \<infinity>) \<or> (b = -\<infinity> \<and> a \<noteq> \<infinity>)"
 10.2093 -  by (cases rule: extreal2_cases[of a b]) auto
 10.2094 -
 10.2095 -lemma extreal_less_divide_pos:
 10.2096 -  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y < z / x \<longleftrightarrow> x * y < z"
 10.2097 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.2098 -
 10.2099 -lemma extreal_divide_less_pos:
 10.2100 -  "x > 0 \<Longrightarrow> x \<noteq> \<infinity> \<Longrightarrow> y / x < z \<longleftrightarrow> y < x * z"
 10.2101 -  by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps)
 10.2102 -
 10.2103 -lemma extreal_divide_eq:
 10.2104 -  "b \<noteq> 0 \<Longrightarrow> \<bar>b\<bar> \<noteq> \<infinity> \<Longrightarrow> a / b = c \<longleftrightarrow> a = b * c"
 10.2105 -  by (cases rule: extreal3_cases[of a b c])
 10.2106 -     (simp_all add: field_simps)
 10.2107 -
 10.2108 -lemma extreal_inverse_not_MInfty[simp]: "inverse a \<noteq> -\<infinity>"
 10.2109 -  by (cases a) auto
 10.2110 -
 10.2111 -lemma extreal_mult_m1[simp]: "x * extreal (-1) = -x"
 10.2112 -  by (cases x) auto
 10.2113 -
 10.2114 -lemma extreal_LimI_finite:
 10.2115 -  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
 10.2116 -  assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r"
 10.2117 -  shows "u ----> x"
 10.2118 -proof (rule topological_tendstoI, unfold eventually_sequentially)
 10.2119 -  obtain rx where rx_def: "x=extreal rx" using assms by (cases x) auto
 10.2120 -  fix S assume "open S" "x : S"
 10.2121 -  then have "open (extreal -` S)" unfolding open_extreal_def by auto
 10.2122 -  with `x \<in> S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> extreal y \<in> S"
 10.2123 -    unfolding open_real_def rx_def by auto
 10.2124 -  then obtain n where
 10.2125 -    upper: "!!N. n <= N ==> u N < x + extreal r" and
 10.2126 -    lower: "!!N. n <= N ==> x < u N + extreal r" using assms(2)[of "extreal r"] by auto
 10.2127 -  show "EX N. ALL n>=N. u n : S"
 10.2128 -  proof (safe intro!: exI[of _ n])
 10.2129 -    fix N assume "n <= N"
 10.2130 -    from upper[OF this] lower[OF this] assms `0 < r`
 10.2131 -    have "u N ~: {\<infinity>,(-\<infinity>)}" by auto
 10.2132 -    from this obtain ra where ra_def: "(u N) = extreal ra" by (cases "u N") auto
 10.2133 -    hence "rx < ra + r" and "ra < rx + r"
 10.2134 -       using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto
 10.2135 -    hence "dist (real (u N)) rx < r"
 10.2136 -      using rx_def ra_def
 10.2137 -      by (auto simp: dist_real_def abs_diff_less_iff field_simps)
 10.2138 -    from dist[OF this] show "u N : S" using `u N  ~: {\<infinity>, -\<infinity>}`
 10.2139 -      by (auto simp: extreal_real split: split_if_asm)
 10.2140 -  qed
 10.2141 -qed
 10.2142 -
 10.2143 -lemma extreal_LimI_finite_iff:
 10.2144 -  assumes "\<bar>x\<bar> \<noteq> \<infinity>"
 10.2145 -  shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))"
 10.2146 -  (is "?lhs <-> ?rhs")
 10.2147 -proof
 10.2148 -  assume lim: "u ----> x"
 10.2149 -  { fix r assume "(r::extreal)>0"
 10.2150 -    from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}"
 10.2151 -       apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"])
 10.2152 -       using lim extreal_between[of x r] assms `r>0` by auto
 10.2153 -    hence "EX N. ALL n>=N. u n < x + r & x < u n + r"
 10.2154 -      using extreal_minus_less[of r x] by (cases r) auto
 10.2155 -  } then show "?rhs" by auto
 10.2156 -next
 10.2157 -  assume ?rhs then show "u ----> x"
 10.2158 -    using extreal_LimI_finite[of x] assms by auto
 10.2159 -qed
 10.2160 -
 10.2161 -
 10.2162 -subsubsection {* @{text Liminf} and @{text Limsup} *}
 10.2163 -
 10.2164 -definition
 10.2165 -  "Liminf net f = (GREATEST l. \<forall>y<l. eventually (\<lambda>x. y < f x) net)"
 10.2166 -
 10.2167 -definition
 10.2168 -  "Limsup net f = (LEAST l. \<forall>y>l. eventually (\<lambda>x. f x < y) net)"
 10.2169 -
 10.2170 -lemma Liminf_Sup:
 10.2171 -  fixes f :: "'a => 'b::{complete_lattice, linorder}"
 10.2172 -  shows "Liminf net f = Sup {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net}"
 10.2173 -  by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def)
 10.2174 -
 10.2175 -lemma Limsup_Inf:
 10.2176 -  fixes f :: "'a => 'b::{complete_lattice, linorder}"
 10.2177 -  shows "Limsup net f = Inf {l. \<forall>y>l. eventually (\<lambda>x. f x < y) net}"
 10.2178 -  by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def)
 10.2179 -
 10.2180 -lemma extreal_SupI:
 10.2181 -  fixes x :: extreal
 10.2182 -  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
 10.2183 -  assumes "\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y"
 10.2184 -  shows "Sup A = x"
 10.2185 -  unfolding Sup_extreal_def
 10.2186 -  using assms by (auto intro!: Least_equality)
 10.2187 -
 10.2188 -lemma extreal_InfI:
 10.2189 -  fixes x :: extreal
 10.2190 -  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> i"
 10.2191 -  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x"
 10.2192 -  shows "Inf A = x"
 10.2193 -  unfolding Inf_extreal_def
 10.2194 -  using assms by (auto intro!: Greatest_equality)
 10.2195 -
 10.2196 -lemma Limsup_const:
 10.2197 -  fixes c :: "'a::{complete_lattice, linorder}"
 10.2198 -  assumes ntriv: "\<not> trivial_limit net"
 10.2199 -  shows "Limsup net (\<lambda>x. c) = c"
 10.2200 -  unfolding Limsup_Inf
 10.2201 -proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower)
 10.2202 -  fix x assume *: "\<forall>y>x. eventually (\<lambda>_. c < y) net"
 10.2203 -  show "c \<le> x"
 10.2204 -  proof (rule ccontr)
 10.2205 -    assume "\<not> c \<le> x" then have "x < c" by auto
 10.2206 -    then show False using ntriv * by (auto simp: trivial_limit_def)
 10.2207 -  qed
 10.2208 -qed auto
 10.2209 -
 10.2210 -lemma Liminf_const:
 10.2211 -  fixes c :: "'a::{complete_lattice, linorder}"
 10.2212 -  assumes ntriv: "\<not> trivial_limit net"
 10.2213 -  shows "Liminf net (\<lambda>x. c) = c"
 10.2214 -  unfolding Liminf_Sup
 10.2215 -proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper)
 10.2216 -  fix x assume *: "\<forall>y<x. eventually (\<lambda>_. y < c) net"
 10.2217 -  show "x \<le> c"
 10.2218 -  proof (rule ccontr)
 10.2219 -    assume "\<not> x \<le> c" then have "c < x" by auto
 10.2220 -    then show False using ntriv * by (auto simp: trivial_limit_def)
 10.2221 -  qed
 10.2222 -qed auto
 10.2223 -
 10.2224 -lemma mono_set:
 10.2225 -  fixes S :: "('a::order) set"
 10.2226 -  shows "mono S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)"
 10.2227 -  by (auto simp: mono_def mem_def)
 10.2228 -
 10.2229 -lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto
 10.2230 -lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto
 10.2231 -lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto
 10.2232 -lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto
 10.2233 -
 10.2234 -lemma mono_set_iff:
 10.2235 -  fixes S :: "'a::{linorder,complete_lattice} set"
 10.2236 -  defines "a \<equiv> Inf S"
 10.2237 -  shows "mono S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
 10.2238 -proof
 10.2239 -  assume "mono S"
 10.2240 -  then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set)
 10.2241 -  show ?c
 10.2242 -  proof cases
 10.2243 -    assume "a \<in> S"
 10.2244 -    show ?c
 10.2245 -      using mono[OF _ `a \<in> S`]
 10.2246 -      by (auto intro: complete_lattice_class.Inf_lower simp: a_def)
 10.2247 -  next
 10.2248 -    assume "a \<notin> S"
 10.2249 -    have "S = {a <..}"
 10.2250 -    proof safe
 10.2251 -      fix x assume "x \<in> S"
 10.2252 -      then have "a \<le> x" unfolding a_def by (rule complete_lattice_class.Inf_lower)
 10.2253 -      then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto
 10.2254 -    next
 10.2255 -      fix x assume "a < x"
 10.2256 -      then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff ..
 10.2257 -      with mono[of y x] show "x \<in> S" by auto
 10.2258 -    qed
 10.2259 -    then show ?c ..
 10.2260 -  qed
 10.2261 -qed auto
 10.2262 -
 10.2263 -lemma lim_imp_Liminf:
 10.2264 -  fixes f :: "'a \<Rightarrow> extreal"
 10.2265 -  assumes ntriv: "\<not> trivial_limit net"
 10.2266 -  assumes lim: "(f ---> f0) net"
 10.2267 -  shows "Liminf net f = f0"
 10.2268 -  unfolding Liminf_Sup
 10.2269 -proof (safe intro!: extreal_SupI)
 10.2270 -  fix y assume *: "\<forall>y'<y. eventually (\<lambda>x. y' < f x) net"
 10.2271 -  show "y \<le> f0"
 10.2272 -  proof (rule extreal_le_extreal)
 10.2273 -    fix B assume "B < y"
 10.2274 -    { assume "f0 < B"
 10.2275 -      then have "eventually (\<lambda>x. f x < B \<and> B < f x) net"
 10.2276 -         using topological_tendstoD[OF lim, of "{..<B}"] *[rule_format, OF `B < y`]
 10.2277 -         by (auto intro: eventually_conj)
 10.2278 -      also have "(\<lambda>x. f x < B \<and> B < f x) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
 10.2279 -      finally have False using ntriv[unfolded trivial_limit_def] by auto
 10.2280 -    } then show "B \<le> f0" by (metis linorder_le_less_linear)
 10.2281 -  qed
 10.2282 -next
 10.2283 -  fix y assume *: "\<forall>z. z \<in> {l. \<forall>y<l. eventually (\<lambda>x. y < f x) net} \<longrightarrow> z \<le> y"
 10.2284 -  show "f0 \<le> y"
 10.2285 -  proof (safe intro!: *[rule_format])
 10.2286 -    fix y assume "y < f0" then show "eventually (\<lambda>x. y < f x) net"
 10.2287 -      using lim[THEN topological_tendstoD, of "{y <..}"] by auto
 10.2288 -  qed
 10.2289 -qed
 10.2290 -
 10.2291 -lemma extreal_Liminf_le_Limsup:
 10.2292 -  fixes f :: "'a \<Rightarrow> extreal"
 10.2293 -  assumes ntriv: "\<not> trivial_limit net"
 10.2294 -  shows "Liminf net f \<le> Limsup net f"
 10.2295 -  unfolding Limsup_Inf Liminf_Sup
 10.2296 -proof (safe intro!: complete_lattice_class.Inf_greatest  complete_lattice_class.Sup_least)
 10.2297 -  fix u v assume *: "\<forall>y<u. eventually (\<lambda>x. y < f x) net" "\<forall>y>v. eventually (\<lambda>x. f x < y) net"
 10.2298 -  show "u \<le> v"
 10.2299 -  proof (rule ccontr)
 10.2300 -    assume "\<not> u \<le> v"
 10.2301 -    then obtain t where "t < u" "v < t"
 10.2302 -      using extreal_dense[of v u] by (auto simp: not_le)
 10.2303 -    then have "eventually (\<lambda>x. t < f x \<and> f x < t) net"
 10.2304 -      using * by (auto intro: eventually_conj)
 10.2305 -    also have "(\<lambda>x. t < f x \<and> f x < t) = (\<lambda>x. False)" by (auto simp: fun_eq_iff)
 10.2306 -    finally show False using ntriv by (auto simp: trivial_limit_def)
 10.2307 -  qed
 10.2308 -qed
 10.2309 -
 10.2310 -lemma Liminf_mono:
 10.2311 -  fixes f g :: "'a => extreal"
 10.2312 -  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
 10.2313 -  shows "Liminf net f \<le> Liminf net g"
 10.2314 -  unfolding Liminf_Sup
 10.2315 -proof (safe intro!: Sup_mono bexI)
 10.2316 -  fix a y assume "\<forall>y<a. eventually (\<lambda>x. y < f x) net" and "y < a"
 10.2317 -  then have "eventually (\<lambda>x. y < f x) net" by auto
 10.2318 -  then show "eventually (\<lambda>x. y < g x) net"
 10.2319 -    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
 10.2320 -qed simp
 10.2321 -
 10.2322 -lemma Liminf_eq:
 10.2323 -  fixes f g :: "'a \<Rightarrow> extreal"
 10.2324 -  assumes "eventually (\<lambda>x. f x = g x) net"
 10.2325 -  shows "Liminf net f = Liminf net g"
 10.2326 -  by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
 10.2327 -
 10.2328 -lemma Liminf_mono_all:
 10.2329 -  fixes f g :: "'a \<Rightarrow> extreal"
 10.2330 -  assumes "\<And>x. f x \<le> g x"
 10.2331 -  shows "Liminf net f \<le> Liminf net g"
 10.2332 -  using assms by (intro Liminf_mono always_eventually) auto
 10.2333 -
 10.2334 -lemma Limsup_mono:
 10.2335 -  fixes f g :: "'a \<Rightarrow> extreal"
 10.2336 -  assumes ev: "eventually (\<lambda>x. f x \<le> g x) net"
 10.2337 -  shows "Limsup net f \<le> Limsup net g"
 10.2338 -  unfolding Limsup_Inf
 10.2339 -proof (safe intro!: Inf_mono bexI)
 10.2340 -  fix a y assume "\<forall>y>a. eventually (\<lambda>x. g x < y) net" and "a < y"
 10.2341 -  then have "eventually (\<lambda>x. g x < y) net" by auto
 10.2342 -  then show "eventually (\<lambda>x. f x < y) net"
 10.2343 -    by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto)
 10.2344 -qed simp
 10.2345 -
 10.2346 -lemma Limsup_mono_all:
 10.2347 -  fixes f g :: "'a \<Rightarrow> extreal"
 10.2348 -  assumes "\<And>x. f x \<le> g x"
 10.2349 -  shows "Limsup net f \<le> Limsup net g"
 10.2350 -  using assms by (intro Limsup_mono always_eventually) auto
 10.2351 -
 10.2352 -lemma Limsup_eq:
 10.2353 -  fixes f g :: "'a \<Rightarrow> extreal"
 10.2354 -  assumes "eventually (\<lambda>x. f x = g x) net"
 10.2355 -  shows "Limsup net f = Limsup net g"
 10.2356 -  by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
 10.2357 -
 10.2358 -abbreviation "liminf \<equiv> Liminf sequentially"
 10.2359 -
 10.2360 -abbreviation "limsup \<equiv> Limsup sequentially"
 10.2361 -
 10.2362 -lemma (in complete_lattice) less_INFD:
 10.2363 -  assumes "y < INFI A f"" i \<in> A" shows "y < f i"
 10.2364 -proof -
 10.2365 -  note `y < INFI A f`
 10.2366 -  also have "INFI A f \<le> f i" using `i \<in> A` by (rule INF_leI)
 10.2367 -  finally show "y < f i" .
 10.2368 -qed
 10.2369 -
 10.2370 -lemma liminf_SUPR_INFI:
 10.2371 -  fixes f :: "nat \<Rightarrow> extreal"
 10.2372 -  shows "liminf f = (SUP n. INF m:{n..}. f m)"
 10.2373 -  unfolding Liminf_Sup eventually_sequentially
 10.2374 -proof (safe intro!: antisym complete_lattice_class.Sup_least)
 10.2375 -  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)"
 10.2376 -  proof (rule extreal_le_extreal)
 10.2377 -    fix y assume "y < x"
 10.2378 -    with * obtain N where "\<And>n. N \<le> n \<Longrightarrow> y < f n" by auto
 10.2379 -    then have "y \<le> (INF m:{N..}. f m)" by (force simp: le_INF_iff)
 10.2380 -    also have "\<dots> \<le> (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto
 10.2381 -    finally show "y \<le> (SUP n. INF m:{n..}. f m)" .
 10.2382 -  qed
 10.2383 -next
 10.2384 -  show "(SUP n. INF m:{n..}. f m) \<le> Sup {l. \<forall>y<l. \<exists>N. \<forall>n\<ge>N. y < f n}"
 10.2385 -  proof (unfold SUPR_def, safe intro!: Sup_mono bexI)
 10.2386 -    fix y n assume "y < INFI {n..} f"
 10.2387 -    from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. y < f n" by (intro exI[of _ n]) auto
 10.2388 -  qed (rule order_refl)
 10.2389 -qed
 10.2390 -
 10.2391 -lemma tail_same_limsup:
 10.2392 -  fixes X Y :: "nat => extreal"
 10.2393 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
 10.2394 -  shows "limsup X = limsup Y"
 10.2395 -  using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto
 10.2396 -
 10.2397 -lemma tail_same_liminf:
 10.2398 -  fixes X Y :: "nat => extreal"
 10.2399 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n = Y n"
 10.2400 -  shows "liminf X = liminf Y"
 10.2401 -  using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto
 10.2402 -
 10.2403 -lemma liminf_mono:
 10.2404 -  fixes X Y :: "nat \<Rightarrow> extreal"
 10.2405 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
 10.2406 -  shows "liminf X \<le> liminf Y"
 10.2407 -  using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto
 10.2408 -
 10.2409 -lemma limsup_mono:
 10.2410 -  fixes X Y :: "nat => extreal"
 10.2411 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
 10.2412 -  shows "limsup X \<le> limsup Y"
 10.2413 -  using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto
 10.2414 -
 10.2415 -declare trivial_limit_sequentially[simp]
 10.2416 -
 10.2417 -lemma
 10.2418 -  fixes X :: "nat \<Rightarrow> extreal"
 10.2419 -  shows extreal_incseq_uminus[simp]: "incseq (\<lambda>i. - X i) = decseq X"
 10.2420 -    and extreal_decseq_uminus[simp]: "decseq (\<lambda>i. - X i) = incseq X"
 10.2421 -  unfolding incseq_def decseq_def by auto
 10.2422 -
 10.2423 -lemma liminf_bounded:
 10.2424 -  fixes X Y :: "nat \<Rightarrow> extreal"
 10.2425 -  assumes "\<And>n. N \<le> n \<Longrightarrow> C \<le> X n"
 10.2426 -  shows "C \<le> liminf X"
 10.2427 -  using liminf_mono[of N "\<lambda>n. C" X] assms Liminf_const[of sequentially C] by simp
 10.2428 -
 10.2429 -lemma limsup_bounded:
 10.2430 -  fixes X Y :: "nat => extreal"
 10.2431 -  assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= C"
 10.2432 -  shows "limsup X \<le> C"
 10.2433 -  using limsup_mono[of N X "\<lambda>n. C"] assms Limsup_const[of sequentially C] by simp
 10.2434 -
 10.2435 -lemma liminf_bounded_iff:
 10.2436 -  fixes x :: "nat \<Rightarrow> extreal"
 10.2437 -  shows "C \<le> liminf x \<longleftrightarrow> (\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n)" (is "?lhs <-> ?rhs")
 10.2438 -proof safe
 10.2439 -  fix B assume "B < C" "C \<le> liminf x"
 10.2440 -  then have "B < liminf x" by auto
 10.2441 -  then obtain N where "B < (INF m:{N..}. x m)"
 10.2442 -    unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto
 10.2443 -  from less_INFD[OF this] show "\<exists>N. \<forall>n\<ge>N. B < x n" by auto
 10.2444 -next
 10.2445 -  assume *: "\<forall>B<C. \<exists>N. \<forall>n\<ge>N. B < x n"
 10.2446 -  { fix B assume "B<C"
 10.2447 -    then obtain N where "\<forall>n\<ge>N. B < x n" using `?rhs` by auto
 10.2448 -    hence "B \<le> (INF m:{N..}. x m)" by (intro le_INFI) auto
 10.2449 -    also have "... \<le> liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp
 10.2450 -    finally have "B \<le> liminf x" .
 10.2451 -  } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear)
 10.2452 -qed
 10.2453 -
 10.2454 -lemma liminf_subseq_mono:
 10.2455 -  fixes X :: "nat \<Rightarrow> extreal"
 10.2456 -  assumes "subseq r"
 10.2457 -  shows "liminf X \<le> liminf (X \<circ> r) "
 10.2458 -proof-
 10.2459 -  have "\<And>n. (INF m:{n..}. X m) \<le> (INF m:{n..}. (X \<circ> r) m)"
 10.2460 -  proof (safe intro!: INF_mono)
 10.2461 -    fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m"
 10.2462 -      using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto
 10.2463 -  qed
 10.2464 -  then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def)
 10.2465 -qed
 10.2466 -
 10.2467 -lemma extreal_real': assumes "\<bar>x\<bar> \<noteq> \<infinity>" shows "extreal (real x) = x"
 10.2468 -  using assms by auto
 10.2469 -
 10.2470 -lemma extreal_le_extreal_bounded:
 10.2471 -  fixes x y z :: extreal
 10.2472 -  assumes "z \<le> y"
 10.2473 -  assumes *: "\<And>B. z < B \<Longrightarrow> B < x \<Longrightarrow> B \<le> y"
 10.2474 -  shows "x \<le> y"
 10.2475 -proof (rule extreal_le_extreal)
 10.2476 -  fix B assume "B < x"
 10.2477 -  show "B \<le> y"
 10.2478 -  proof cases
 10.2479 -    assume "z < B" from *[OF this `B < x`] show "B \<le> y" .
 10.2480 -  next
 10.2481 -    assume "\<not> z < B" with `z \<le> y` show "B \<le> y" by auto
 10.2482 -  qed
 10.2483 -qed
 10.2484 -
 10.2485 -lemma fixes x y :: extreal
 10.2486 -  shows Sup_atMost[simp]: "Sup {.. y} = y"
 10.2487 -    and Sup_lessThan[simp]: "Sup {..< y} = y"
 10.2488 -    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
 10.2489 -    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
 10.2490 -    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
 10.2491 -  by (auto simp: Sup_extreal_def intro!: Least_equality
 10.2492 -           intro: extreal_le_extreal extreal_le_extreal_bounded[of x])
 10.2493 -
 10.2494 -lemma Sup_greaterThanlessThan[simp]:
 10.2495 -  fixes x y :: extreal assumes "x < y" shows "Sup { x <..< y} = y"
 10.2496 -  unfolding Sup_extreal_def
 10.2497 -proof (intro Least_equality extreal_le_extreal_bounded[of _ _ y])
 10.2498 -  fix z assume z: "\<forall>u\<in>{x<..<y}. u \<le> z"
 10.2499 -  from extreal_dense[OF `x < y`] guess w .. note w = this
 10.2500 -  with z[THEN bspec, of w] show "x \<le> z" by auto
 10.2501 -qed auto
 10.2502 -
 10.2503 -lemma real_extreal_id: "real o extreal = id"
 10.2504 -proof-
 10.2505 -{ fix x have "(real o extreal) x = id x" by auto }
 10.2506 -from this show ?thesis using ext by blast
 10.2507 -qed
 10.2508 -
 10.2509 -lemma open_image_extreal: "open(UNIV-{\<infinity>,(-\<infinity>)})"
 10.2510 -by (metis range_extreal open_extreal open_UNIV)
 10.2511 -
 10.2512 -lemma extreal_le_distrib:
 10.2513 -  fixes a b c :: extreal shows "c * (a + b) \<le> c * a + c * b"
 10.2514 -  by (cases rule: extreal3_cases[of a b c])
 10.2515 -     (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
 10.2516 -
 10.2517 -lemma extreal_pos_distrib:
 10.2518 -  fixes a b c :: extreal assumes "0 \<le> c" "c \<noteq> \<infinity>" shows "c * (a + b) = c * a + c * b"
 10.2519 -  using assms by (cases rule: extreal3_cases[of a b c])
 10.2520 -                 (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff)
 10.2521 -
 10.2522 -lemma extreal_pos_le_distrib:
 10.2523 -fixes a b c :: extreal
 10.2524 -assumes "c>=0"
 10.2525 -shows "c * (a + b) <= c * a + c * b"
 10.2526 -  using assms by (cases rule: extreal3_cases[of a b c])
 10.2527 -                 (auto simp add: field_simps)
 10.2528 -
 10.2529 -lemma extreal_max_mono:
 10.2530 -  "[| (a::extreal) <= b; c <= d |] ==> max a c <= max b d"
 10.2531 -  by (metis sup_extreal_def sup_mono)
 10.2532 -
 10.2533 -
 10.2534 -lemma extreal_max_least:
 10.2535 -  "[| (a::extreal) <= x; c <= x |] ==> max a c <= x"
 10.2536 -  by (metis sup_extreal_def sup_least)
 10.2537 -
 10.2538 -end
    11.1 --- a/src/HOL/Library/Library.thy	Wed Jul 20 10:11:08 2011 +0200
    11.2 +++ b/src/HOL/Library/Library.thy	Wed Jul 20 10:48:00 2011 +0200
    11.3 @@ -15,6 +15,7 @@
    11.4    Diagonalize
    11.5    Dlist_Cset
    11.6    Eval_Witness
    11.7 +  Extended_Nat
    11.8    Float
    11.9    Formal_Power_Series
   11.10    Fraction_Field
   11.11 @@ -35,7 +36,6 @@
   11.12    Monad_Syntax
   11.13    More_List
   11.14    Multiset
   11.15 -  Nat_Infinity
   11.16    Nested_Environment
   11.17    Numeral_Type
   11.18    OptionalSugar
    12.1 --- a/src/HOL/Library/Nat_Infinity.thy	Wed Jul 20 10:11:08 2011 +0200
    12.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    12.3 @@ -1,551 +0,0 @@
    12.4 -(*  Title:      HOL/Library/Nat_Infinity.thy
    12.5 -    Author:     David von Oheimb, TU Muenchen;  Florian Haftmann, TU Muenchen
    12.6 -    Contributions: David Trachtenherz, TU Muenchen
    12.7 -*)
    12.8 -
    12.9 -header {* Natural numbers with infinity *}
   12.10 -
   12.11 -theory Nat_Infinity
   12.12 -imports Main
   12.13 -begin
   12.14 -
   12.15 -subsection {* Type definition *}
   12.16 -
   12.17 -text {*
   12.18 -  We extend the standard natural numbers by a special value indicating
   12.19 -  infinity.
   12.20 -*}
   12.21 -
   12.22 -datatype inat = Fin nat | Infty
   12.23 -
   12.24 -notation (xsymbols)
   12.25 -  Infty  ("\<infinity>")
   12.26 -
   12.27 -notation (HTML output)
   12.28 -  Infty  ("\<infinity>")
   12.29 -
   12.30 -
   12.31 -lemma not_Infty_eq[iff]: "(x ~= Infty) = (EX i. x = Fin i)"
   12.32 -by (cases x) auto
   12.33 -
   12.34 -lemma not_Fin_eq [iff]: "(ALL y. x ~= Fin y) = (x = Infty)"
   12.35 -by (cases x) auto
   12.36 -
   12.37 -
   12.38 -primrec the_Fin :: "inat \<Rightarrow> nat"
   12.39 -where "the_Fin (Fin n) = n"
   12.40 -
   12.41 -
   12.42 -subsection {* Constructors and numbers *}
   12.43 -
   12.44 -instantiation inat :: "{zero, one, number}"
   12.45 -begin
   12.46 -
   12.47 -definition
   12.48 -  "0 = Fin 0"
   12.49 -
   12.50 -definition
   12.51 -  [code_unfold]: "1 = Fin 1"
   12.52 -
   12.53 -definition
   12.54 -  [code_unfold, code del]: "number_of k = Fin (number_of k)"
   12.55 -
   12.56 -instance ..
   12.57 -
   12.58 -end
   12.59 -
   12.60 -definition iSuc :: "inat \<Rightarrow> inat" where
   12.61 -  "iSuc i = (case i of Fin n \<Rightarrow> Fin (Suc n) | \<infinity> \<Rightarrow> \<infinity>)"
   12.62 -
   12.63 -lemma Fin_0: "Fin 0 = 0"
   12.64 -  by (simp add: zero_inat_def)
   12.65 -
   12.66 -lemma Fin_1: "Fin 1 = 1"
   12.67 -  by (simp add: one_inat_def)
   12.68 -
   12.69 -lemma Fin_number: "Fin (number_of k) = number_of k"
   12.70 -  by (simp add: number_of_inat_def)
   12.71 -
   12.72 -lemma one_iSuc: "1 = iSuc 0"
   12.73 -  by (simp add: zero_inat_def one_inat_def iSuc_def)
   12.74 -
   12.75 -lemma Infty_ne_i0 [simp]: "\<infinity> \<noteq> 0"
   12.76 -  by (simp add: zero_inat_def)
   12.77 -
   12.78 -lemma i0_ne_Infty [simp]: "0 \<noteq> \<infinity>"
   12.79 -  by (simp add: zero_inat_def)
   12.80 -
   12.81 -lemma zero_inat_eq [simp]:
   12.82 -  "number_of k = (0\<Colon>inat) \<longleftrightarrow> number_of k = (0\<Colon>nat)"
   12.83 -  "(0\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (0\<Colon>nat)"
   12.84 -  unfolding zero_inat_def number_of_inat_def by simp_all
   12.85 -
   12.86 -lemma one_inat_eq [simp]:
   12.87 -  "number_of k = (1\<Colon>inat) \<longleftrightarrow> number_of k = (1\<Colon>nat)"
   12.88 -  "(1\<Colon>inat) = number_of k \<longleftrightarrow> number_of k = (1\<Colon>nat)"
   12.89 -  unfolding one_inat_def number_of_inat_def by simp_all
   12.90 -
   12.91 -lemma zero_one_inat_neq [simp]:
   12.92 -  "\<not> 0 = (1\<Colon>inat)"
   12.93 -  "\<not> 1 = (0\<Colon>inat)"
   12.94 -  unfolding zero_inat_def one_inat_def by simp_all
   12.95 -
   12.96 -lemma Infty_ne_i1 [simp]: "\<infinity> \<noteq> 1"
   12.97 -  by (simp add: one_inat_def)
   12.98 -
   12.99 -lemma i1_ne_Infty [simp]: "1 \<noteq> \<infinity>"
  12.100 -  by (simp add: one_inat_def)
  12.101 -
  12.102 -lemma Infty_ne_number [simp]: "\<infinity> \<noteq> number_of k"
  12.103 -  by (simp add: number_of_inat_def)
  12.104 -
  12.105 -lemma number_ne_Infty [simp]: "number_of k \<noteq> \<infinity>"
  12.106 -  by (simp add: number_of_inat_def)
  12.107 -
  12.108 -lemma iSuc_Fin: "iSuc (Fin n) = Fin (Suc n)"
  12.109 -  by (simp add: iSuc_def)
  12.110 -
  12.111 -lemma iSuc_number_of: "iSuc (number_of k) = Fin (Suc (number_of k))"
  12.112 -  by (simp add: iSuc_Fin number_of_inat_def)
  12.113 -
  12.114 -lemma iSuc_Infty [simp]: "iSuc \<infinity> = \<infinity>"
  12.115 -  by (simp add: iSuc_def)
  12.116 -
  12.117 -lemma iSuc_ne_0 [simp]: "iSuc n \<noteq> 0"
  12.118 -  by (simp add: iSuc_def zero_inat_def split: inat.splits)
  12.119 -
  12.120 -lemma zero_ne_iSuc [simp]: "0 \<noteq> iSuc n"
  12.121 -  by (rule iSuc_ne_0 [symmetric])
  12.122 -
  12.123 -lemma iSuc_inject [simp]: "iSuc m = iSuc n \<longleftrightarrow> m = n"
  12.124 -  by (simp add: iSuc_def split: inat.splits)
  12.125 -
  12.126 -lemma number_of_inat_inject [simp]:
  12.127 -  "(number_of k \<Colon> inat) = number_of l \<longleftrightarrow> (number_of k \<Colon> nat) = number_of l"
  12.128 -  by (simp add: number_of_inat_def)
  12.129 -
  12.130 -
  12.131 -subsection {* Addition *}
  12.132 -
  12.133 -instantiation inat :: comm_monoid_add
  12.134 -begin
  12.135 -
  12.136 -definition [nitpick_simp]:
  12.137 -  "m + n = (case m of \<infinity> \<Rightarrow> \<infinity> | Fin m \<Rightarrow> (case n of \<infinity> \<Rightarrow> \<infinity> | Fin n \<Rightarrow> Fin (m + n)))"
  12.138 -
  12.139 -lemma plus_inat_simps [simp, code]:
  12.140 -  "Fin m + Fin n = Fin (m + n)"
  12.141 -  "\<infinity> + q = \<infinity>"
  12.142 -  "q + \<infinity> = \<infinity>"
  12.143 -  by (simp_all add: plus_inat_def split: inat.splits)
  12.144 -
  12.145 -instance proof
  12.146 -  fix n m q :: inat
  12.147 -  show "n + m + q = n + (m + q)"
  12.148 -    by (cases n, auto, cases m, auto, cases q, auto)
  12.149 -  show "n + m = m + n"
  12.150 -    by (cases n, auto, cases m, auto)
  12.151 -  show "0 + n = n"
  12.152 -    by (cases n) (simp_all add: zero_inat_def)
  12.153 -qed
  12.154 -
  12.155 -end
  12.156 -
  12.157 -lemma plus_inat_0 [simp]:
  12.158 -  "0 + (q\<Colon>inat) = q"
  12.159 -  "(q\<Colon>inat) + 0 = q"
  12.160 -  by (simp_all add: plus_inat_def zero_inat_def split: inat.splits)
  12.161 -
  12.162 -lemma plus_inat_number [simp]:
  12.163 -  "(number_of k \<Colon> inat) + number_of l = (if k < Int.Pls then number_of l
  12.164 -    else if l < Int.Pls then number_of k else number_of (k + l))"
  12.165 -  unfolding number_of_inat_def plus_inat_simps nat_arith(1) if_distrib [symmetric, of _ Fin] ..
  12.166 -
  12.167 -lemma iSuc_number [simp]:
  12.168 -  "iSuc (number_of k) = (if neg (number_of k \<Colon> int) then 1 else number_of (Int.succ k))"
  12.169 -  unfolding iSuc_number_of
  12.170 -  unfolding one_inat_def number_of_inat_def Suc_nat_number_of if_distrib [symmetric] ..
  12.171 -
  12.172 -lemma iSuc_plus_1:
  12.173 -  "iSuc n = n + 1"
  12.174 -  by (cases n) (simp_all add: iSuc_Fin one_inat_def)
  12.175 -  
  12.176 -lemma plus_1_iSuc:
  12.177 -  "1 + q = iSuc q"
  12.178 -  "q + 1 = iSuc q"
  12.179 -by (simp_all add: iSuc_plus_1 add_ac)
  12.180 -
  12.181 -lemma iadd_Suc: "iSuc m + n = iSuc (m + n)"
  12.182 -by (simp_all add: iSuc_plus_1 add_ac)
  12.183 -
  12.184 -lemma iadd_Suc_right: "m + iSuc n = iSuc (m + n)"
  12.185 -by (simp only: add_commute[of m] iadd_Suc)
  12.186 -
  12.187 -lemma iadd_is_0: "(m + n = (0::inat)) = (m = 0 \<and> n = 0)"
  12.188 -by (cases m, cases n, simp_all add: zero_inat_def)
  12.189 -
  12.190 -subsection {* Multiplication *}
  12.191 -
  12.192 -instantiation inat :: comm_semiring_1
  12.193 -begin
  12.194 -
  12.195 -definition times_inat_def [nitpick_simp]:
  12.196 -  "m * n = (case m of \<infinity> \<Rightarrow> if n = 0 then 0 else \<infinity> | Fin m \<Rightarrow>
  12.197 -    (case n of \<infinity> \<Rightarrow> if m = 0 then 0 else \<infinity> | Fin n \<Rightarrow> Fin (m * n)))"
  12.198 -
  12.199 -lemma times_inat_simps [simp, code]:
  12.200 -  "Fin m * Fin n = Fin (m * n)"
  12.201 -  "\<infinity> * \<infinity> = \<infinity>"
  12.202 -  "\<infinity> * Fin n = (if n = 0 then 0 else \<infinity>)"
  12.203 -  "Fin m * \<infinity> = (if m = 0 then 0 else \<infinity>)"
  12.204 -  unfolding times_inat_def zero_inat_def
  12.205 -  by (simp_all split: inat.split)
  12.206 -
  12.207 -instance proof
  12.208 -  fix a b c :: inat
  12.209 -  show "(a * b) * c = a * (b * c)"
  12.210 -    unfolding times_inat_def zero_inat_def
  12.211 -    by (simp split: inat.split)
  12.212 -  show "a * b = b * a"
  12.213 -    unfolding times_inat_def zero_inat_def
  12.214 -    by (simp split: inat.split)
  12.215 -  show "1 * a = a"
  12.216 -    unfolding times_inat_def zero_inat_def one_inat_def
  12.217 -    by (simp split: inat.split)
  12.218 -  show "(a + b) * c = a * c + b * c"
  12.219 -    unfolding times_inat_def zero_inat_def
  12.220 -    by (simp split: inat.split add: left_distrib)
  12.221 -  show "0 * a = 0"
  12.222 -    unfolding times_inat_def zero_inat_def
  12.223 -    by (simp split: inat.split)
  12.224 -  show "a * 0 = 0"
  12.225 -    unfolding times_inat_def zero_inat_def
  12.226 -    by (simp split: inat.split)
  12.227 -  show "(0::inat) \<noteq> 1"
  12.228 -    unfolding zero_inat_def one_inat_def
  12.229 -    by simp
  12.230 -qed
  12.231 -
  12.232 -end
  12.233 -
  12.234 -lemma mult_iSuc: "iSuc m * n = n + m * n"
  12.235 -  unfolding iSuc_plus_1 by (simp add: algebra_simps)
  12.236 -
  12.237 -lemma mult_iSuc_right: "m * iSuc n = m + m * n"
  12.238 -  unfolding iSuc_plus_1 by (simp add: algebra_simps)
  12.239 -
  12.240 -lemma of_nat_eq_Fin: "of_nat n = Fin n"
  12.241 -  apply (induct n)
  12.242 -  apply (simp add: Fin_0)
  12.243 -  apply (simp add: plus_1_iSuc iSuc_Fin)
  12.244 -  done
  12.245 -
  12.246 -instance inat :: number_semiring
  12.247 -proof
  12.248 -  fix n show "number_of (int n) = (of_nat n :: inat)"
  12.249 -    unfolding number_of_inat_def number_of_int of_nat_id of_nat_eq_Fin ..
  12.250 -qed
  12.251 -
  12.252 -instance inat :: semiring_char_0 proof
  12.253 -  have "inj Fin" by (rule injI) simp
  12.254 -  then show "inj (\<lambda>n. of_nat n :: inat)" by (simp add: of_nat_eq_Fin)
  12.255 -qed
  12.256 -
  12.257 -lemma imult_is_0[simp]: "((m::inat) * n = 0) = (m = 0 \<or> n = 0)"
  12.258 -by(auto simp add: times_inat_def zero_inat_def split: inat.split)
  12.259 -
  12.260 -lemma imult_is_Infty: "((a::inat) * b = \<infinity>) = (a = \<infinity> \<and> b \<noteq> 0 \<or> b = \<infinity> \<and> a \<noteq> 0)"
  12.261 -by(auto simp add: times_inat_def zero_inat_def split: inat.split)
  12.262 -
  12.263 -
  12.264 -subsection {* Subtraction *}
  12.265 -
  12.266 -instantiation inat :: minus
  12.267 -begin
  12.268 -
  12.269 -definition diff_inat_def:
  12.270 -"a - b = (case a of (Fin x) \<Rightarrow> (case b of (Fin y) \<Rightarrow> Fin (x - y) | \<infinity> \<Rightarrow> 0)
  12.271 -          | \<infinity> \<Rightarrow> \<infinity>)"
  12.272 -
  12.273 -instance ..
  12.274 -
  12.275 -end
  12.276 -
  12.277 -lemma idiff_Fin_Fin[simp,code]: "Fin a - Fin b = Fin (a - b)"
  12.278 -by(simp add: diff_inat_def)
  12.279 -
  12.280 -lemma idiff_Infty[simp,code]: "\<infinity> - n = \<infinity>"
  12.281 -by(simp add: diff_inat_def)
  12.282 -
  12.283 -lemma idiff_Infty_right[simp,code]: "Fin a - \<infinity> = 0"
  12.284 -by(simp add: diff_inat_def)
  12.285 -
  12.286 -lemma idiff_0[simp]: "(0::inat) - n = 0"
  12.287 -by (cases n, simp_all add: zero_inat_def)
  12.288 -
  12.289 -lemmas idiff_Fin_0[simp] = idiff_0[unfolded zero_inat_def]
  12.290 -
  12.291 -lemma idiff_0_right[simp]: "(n::inat) - 0 = n"
  12.292 -by (cases n) (simp_all add: zero_inat_def)
  12.293 -
  12.294 -lemmas idiff_Fin_0_right[simp] = idiff_0_right[unfolded zero_inat_def]
  12.295 -
  12.296 -lemma idiff_self[simp]: "n \<noteq> \<infinity> \<Longrightarrow> (n::inat) - n = 0"
  12.297 -by(auto simp: zero_inat_def)
  12.298 -
  12.299 -lemma iSuc_minus_iSuc [simp]: "iSuc n - iSuc m = n - m"
  12.300 -by(simp add: iSuc_def split: inat.split)
  12.301 -
  12.302 -lemma iSuc_minus_1 [simp]: "iSuc n - 1 = n"
  12.303 -by(simp add: one_inat_def iSuc_Fin[symmetric] zero_inat_def[symmetric])
  12.304 -
  12.305 -(*lemmas idiff_self_eq_0_Fin = idiff_self_eq_0[unfolded zero_inat_def]*)
  12.306 -
  12.307 -
  12.308 -subsection {* Ordering *}
  12.309 -
  12.310 -instantiation inat :: linordered_ab_semigroup_add
  12.311 -begin
  12.312 -
  12.313 -definition [nitpick_simp]:
  12.314 -  "m \<le> n = (case n of Fin n1 \<Rightarrow> (case m of Fin m1 \<Rightarrow> m1 \<le> n1 | \<infinity> \<Rightarrow> False)
  12.315 -    | \<infinity> \<Rightarrow> True)"
  12.316 -
  12.317 -definition [nitpick_simp]:
  12.318 -  "m < n = (case m of Fin m1 \<Rightarrow> (case n of Fin n1 \<Rightarrow> m1 < n1 | \<infinity> \<Rightarrow> True)
  12.319 -    | \<infinity> \<Rightarrow> False)"
  12.320 -
  12.321 -lemma inat_ord_simps [simp]:
  12.322 -  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
  12.323 -  "Fin m < Fin n \<longleftrightarrow> m < n"
  12.324 -  "q \<le> \<infinity>"
  12.325 -  "q < \<infinity> \<longleftrightarrow> q \<noteq> \<infinity>"
  12.326 -  "\<infinity> \<le> q \<longleftrightarrow> q = \<infinity>"
  12.327 -  "\<infinity> < q \<longleftrightarrow> False"
  12.328 -  by (simp_all add: less_eq_inat_def less_inat_def split: inat.splits)
  12.329 -
  12.330 -lemma inat_ord_code [code]:
  12.331 -  "Fin m \<le> Fin n \<longleftrightarrow> m \<le> n"
  12.332 -  "Fin m < Fin n \<longleftrightarrow> m < n"
  12.333 -  "q \<le> \<infinity> \<longleftrightarrow> True"
  12.334 -  "Fin m < \<infinity> \<longleftrightarrow> True"
  12.335 -  "\<infinity> \<le> Fin n \<longleftrightarrow> False"
  12.336 -  "\<infinity> < q \<longleftrightarrow> False"
  12.337 -  by simp_all
  12.338 -
  12.339 -instance by default
  12.340 -  (auto simp add: less_eq_inat_def less_inat_def plus_inat_def split: inat.splits)
  12.341 -
  12.342 -end
  12.343 -
  12.344 -instance inat :: ordered_comm_semiring
  12.345 -proof
  12.346 -  fix a b c :: inat
  12.347 -  assume "a \<le> b" and "0 \<le> c"
  12.348 -  thus "c * a \<le> c * b"
  12.349 -    unfolding times_inat_def less_eq_inat_def zero_inat_def
  12.350 -    by (simp split: inat.splits)
  12.351 -qed
  12.352 -
  12.353 -lemma inat_ord_number [simp]:
  12.354 -  "(number_of m \<Colon> inat) \<le> number_of n \<longleftrightarrow> (number_of m \<Colon> nat) \<le> number_of n"
  12.355 -  "(number_of m \<Colon> inat) < number_of n \<longleftrightarrow> (number_of m \<Colon> nat) < number_of n"
  12.356 -  by (simp_all add: number_of_inat_def)
  12.357 -
  12.358 -lemma i0_lb [simp]: "(0\<Colon>inat) \<le> n"
  12.359 -  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
  12.360 -
  12.361 -lemma ile0_eq [simp]: "n \<le> (0\<Colon>inat) \<longleftrightarrow> n = 0"
  12.362 -by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
  12.363 -
  12.364 -lemma Infty_ileE [elim!]: "\<infinity> \<le> Fin m \<Longrightarrow> R"
  12.365 -  by (simp add: zero_inat_def less_eq_inat_def split: inat.splits)
  12.366 -
  12.367 -lemma Infty_ilessE [elim!]: "\<infinity> < Fin m \<Longrightarrow> R"
  12.368 -  by simp
  12.369 -
  12.370 -lemma not_iless0 [simp]: "\<not> n < (0\<Colon>inat)"
  12.371 -  by (simp add: zero_inat_def less_inat_def split: inat.splits)
  12.372 -
  12.373 -lemma i0_less [simp]: "(0\<Colon>inat) < n \<longleftrightarrow> n \<noteq> 0"
  12.374 -by (simp add: zero_inat_def less_inat_def split: inat.splits)
  12.375 -
  12.376 -lemma iSuc_ile_mono [simp]: "iSuc n \<le> iSuc m \<longleftrightarrow> n \<le> m"
  12.377 -  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
  12.378 - 
  12.379 -lemma iSuc_mono [simp]: "iSuc n < iSuc m \<longleftrightarrow> n < m"
  12.380 -  by (simp add: iSuc_def less_inat_def split: inat.splits)
  12.381 -
  12.382 -lemma ile_iSuc [simp]: "n \<le> iSuc n"
  12.383 -  by (simp add: iSuc_def less_eq_inat_def split: inat.splits)
  12.384 -
  12.385 -lemma not_iSuc_ilei0 [simp]: "\<not> iSuc n \<le> 0"
  12.386 -  by (simp add: zero_inat_def iSuc_def less_eq_inat_def split: inat.splits)
  12.387 -
  12.388 -lemma i0_iless_iSuc [simp]: "0 < iSuc n"
  12.389 -  by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.splits)
  12.390 -
  12.391 -lemma iless_iSuc0[simp]: "(n < iSuc 0) = (n = 0)"
  12.392 -by (simp add: zero_inat_def iSuc_def less_inat_def split: inat.split)
  12.393 -
  12.394 -lemma ileI1: "m < n \<Longrightarrow> iSuc m \<le> n"
  12.395 -  by (simp add: iSuc_def less_eq_inat_def less_inat_def split: inat.splits)
  12.396 -
  12.397 -lemma Suc_ile_eq: "Fin (Suc m) \<le> n \<longleftrightarrow> Fin m < n"
  12.398 -  by (cases n) auto
  12.399 -
  12.400 -lemma iless_Suc_eq [simp]: "Fin m < iSuc n \<longleftrightarrow> Fin m \<le> n"
  12.401 -  by (auto simp add: iSuc_def less_inat_def split: inat.splits)
  12.402 -
  12.403 -lemma imult_Infty: "(0::inat) < n \<Longrightarrow> \<infinity> * n = \<infinity>"
  12.404 -by (simp add: zero_inat_def less_inat_def split: inat.splits)
  12.405 -
  12.406 -lemma imult_Infty_right: "(0::inat) < n \<Longrightarrow> n * \<infinity> = \<infinity>"
  12.407 -by (simp add: zero_inat_def less_inat_def split: inat.splits)
  12.408 -
  12.409 -lemma inat_0_less_mult_iff: "(0 < (m::inat) * n) = (0 < m \<and> 0 < n)"
  12.410 -by (simp only: i0_less imult_is_0, simp)
  12.411 -
  12.412 -lemma mono_iSuc: "mono iSuc"
  12.413 -by(simp add: mono_def)
  12.414 -
  12.415 -
  12.416 -lemma min_inat_simps [simp]:
  12.417 -  "min (Fin m) (Fin n) = Fin (min m n)"
  12.418 -  "min q 0 = 0"
  12.419 -  "min 0 q = 0"
  12.420 -  "min q \<infinity> = q"
  12.421 -  "min \<infinity> q = q"
  12.422 -  by (auto simp add: min_def)
  12.423 -
  12.424 -lemma max_inat_simps [simp]:
  12.425 -  "max (Fin m) (Fin n) = Fin (max m n)"
  12.426 -  "max q 0 = q"
  12.427 -  "max 0 q = q"
  12.428 -  "max q \<infinity> = \<infinity>"
  12.429 -  "max \<infinity> q = \<infinity>"
  12.430 -  by (simp_all add: max_def)
  12.431 -
  12.432 -lemma Fin_ile: "n \<le> Fin m \<Longrightarrow> \<exists>k. n = Fin k"
  12.433 -  by (cases n) simp_all
  12.434 -
  12.435 -lemma Fin_iless: "n < Fin m \<Longrightarrow> \<exists>k. n = Fin k"
  12.436 -  by (cases n) simp_all
  12.437 -
  12.438 -lemma chain_incr: "\<forall>i. \<exists>j. Y i < Y j ==> \<exists>j. Fin k < Y j"
  12.439 -apply (induct_tac k)
  12.440 - apply (simp (no_asm) only: Fin_0)
  12.441 - apply (fast intro: le_less_trans [OF i0_lb])
  12.442 -apply (erule exE)
  12.443 -apply (drule spec)
  12.444 -apply (erule exE)
  12.445 -apply (drule ileI1)
  12.446 -apply (rule iSuc_Fin [THEN subst])
  12.447 -apply (rule exI)
  12.448 -apply (erule (1) le_less_trans)
  12.449 -done
  12.450 -
  12.451 -instantiation inat :: "{bot, top}"
  12.452 -begin
  12.453 -
  12.454 -definition bot_inat :: inat where
  12.455 -  "bot_inat = 0"
  12.456 -
  12.457 -definition top_inat :: inat where
  12.458 -  "top_inat = \<infinity>"
  12.459 -
  12.460 -instance proof
  12.461 -qed (simp_all add: bot_inat_def top_inat_def)
  12.462 -
  12.463 -end
  12.464 -
  12.465 -lemma finite_Fin_bounded:
  12.466 -  assumes le_fin: "\<And>y. y \<in> A \<Longrightarrow> y \<le> Fin n"
  12.467 -  shows "finite A"
  12.468 -proof (rule finite_subset)
  12.469 -  show "finite (Fin ` {..n})" by blast
  12.470 -
  12.471 -  have "A \<subseteq> {..Fin n}" using le_fin by fastsimp
  12.472 -  also have "\<dots> \<subseteq> Fin ` {..n}"
  12.473 -    by (rule subsetI) (case_tac x, auto)
  12.474 -  finally show "A \<subseteq> Fin ` {..n}" .
  12.475 -qed
  12.476 -
  12.477 -
  12.478 -subsection {* Well-ordering *}
  12.479 -
  12.480 -lemma less_FinE:
  12.481 -  "[| n < Fin m; !!k. n = Fin k ==> k < m ==> P |] ==> P"
  12.482 -by (induct n) auto
  12.483 -
  12.484 -lemma less_InftyE:
  12.485 -  "[| n < Infty; !!k. n = Fin k ==> P |] ==> P"
  12.486 -by (induct n) auto
  12.487 -
  12.488 -lemma inat_less_induct:
  12.489 -  assumes prem: "!!n. \<forall>m::inat. m < n --> P m ==> P n" shows "P n"
  12.490 -proof -
  12.491 -  have P_Fin: "!!k. P (Fin k)"
  12.492 -    apply (rule nat_less_induct)
  12.493 -    apply (rule prem, clarify)
  12.494 -    apply (erule less_FinE, simp)
  12.495 -    done
  12.496 -  show ?thesis
  12.497 -  proof (induct n)
  12.498 -    fix nat
  12.499 -    show "P (Fin nat)" by (rule P_Fin)
  12.500 -  next
  12.501 -    show "P Infty"
  12.502 -      apply (rule prem, clarify)
  12.503 -      apply (erule less_InftyE)
  12.504 -      apply (simp add: P_Fin)
  12.505 -      done
  12.506 -  qed
  12.507 -qed
  12.508 -
  12.509 -instance inat :: wellorder
  12.510 -proof
  12.511 -  fix P and n
  12.512 -  assume hyp: "(\<And>n\<Colon>inat. (\<And>m\<Colon>inat. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
  12.513 -  show "P n" by (blast intro: inat_less_induct hyp)
  12.514 -qed
  12.515 -
  12.516 -subsection {* Complete Lattice *}
  12.517 -
  12.518 -instantiation inat :: complete_lattice
  12.519 -begin
  12.520 -
  12.521 -definition inf_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
  12.522 -  "inf_inat \<equiv> min"
  12.523 -
  12.524 -definition sup_inat :: "inat \<Rightarrow> inat \<Rightarrow> inat" where
  12.525 -  "sup_inat \<equiv> max"
  12.526 -
  12.527 -definition Inf_inat :: "inat set \<Rightarrow> inat" where
  12.528 -  "Inf_inat A \<equiv> if A = {} then \<infinity> else (LEAST x. x \<in> A)"
  12.529 -
  12.530 -definition Sup_inat :: "inat set \<Rightarrow> inat" where
  12.531 -  "Sup_inat A \<equiv> if A = {} then 0
  12.532 -    else if finite A then Max A
  12.533 -                     else \<infinity>"
  12.534 -instance proof
  12.535 -  fix x :: "inat" and A :: "inat set"
  12.536 -  { assume "x \<in> A" then show "Inf A \<le> x"
  12.537 -      unfolding Inf_inat_def by (auto intro: Least_le) }
  12.538 -  { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
  12.539 -      unfolding Inf_inat_def
  12.540 -      by (cases "A = {}") (auto intro: LeastI2_ex) }
  12.541 -  { assume "x \<in> A" then show "x \<le> Sup A"
  12.542 -      unfolding Sup_inat_def by (cases "finite A") auto }
  12.543 -  { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
  12.544 -      unfolding Sup_inat_def using finite_Fin_bounded by auto }
  12.545 -qed (simp_all add: inf_inat_def sup_inat_def)
  12.546 -end
  12.547 -
  12.548 -
  12.549 -subsection {* Traditional theorem names *}
  12.550 -
  12.551 -lemmas inat_defs = zero_inat_def one_inat_def number_of_inat_def iSuc_def
  12.552 -  plus_inat_def less_eq_inat_def less_inat_def
  12.553 -
  12.554 -end
    13.1 --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jul 20 10:11:08 2011 +0200
    13.2 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Wed Jul 20 10:48:00 2011 +0200
    13.3 @@ -8,87 +8,87 @@
    13.4  header {* Limits on the Extended real number line *}
    13.5  
    13.6  theory Extended_Real_Limits
    13.7 -  imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Reals"
    13.8 +  imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Real"
    13.9  begin
   13.10  
   13.11 -lemma continuous_on_extreal[intro, simp]: "continuous_on A extreal"
   13.12 -  unfolding continuous_on_topological open_extreal_def by auto
   13.13 +lemma continuous_on_ereal[intro, simp]: "continuous_on A ereal"
   13.14 +  unfolding continuous_on_topological open_ereal_def by auto
   13.15  
   13.16 -lemma continuous_at_extreal[intro, simp]: "continuous (at x) extreal"
   13.17 +lemma continuous_at_ereal[intro, simp]: "continuous (at x) ereal"
   13.18    using continuous_on_eq_continuous_at[of UNIV] by auto
   13.19  
   13.20 -lemma continuous_within_extreal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) extreal"
   13.21 +lemma continuous_within_ereal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) ereal"
   13.22    using continuous_on_eq_continuous_within[of A] by auto
   13.23  
   13.24 -lemma extreal_open_uminus:
   13.25 -  fixes S :: "extreal set"
   13.26 +lemma ereal_open_uminus:
   13.27 +  fixes S :: "ereal set"
   13.28    assumes "open S"
   13.29    shows "open (uminus ` S)"
   13.30 -  unfolding open_extreal_def
   13.31 +  unfolding open_ereal_def
   13.32  proof (intro conjI impI)
   13.33 -  obtain x y where S: "open (extreal -` S)"
   13.34 -    "\<infinity> \<in> S \<Longrightarrow> {extreal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< extreal y} \<subseteq> S"
   13.35 -    using `open S` unfolding open_extreal_def by auto
   13.36 -  have "extreal -` uminus ` S = uminus ` (extreal -` S)"
   13.37 +  obtain x y where S: "open (ereal -` S)"
   13.38 +    "\<infinity> \<in> S \<Longrightarrow> {ereal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< ereal y} \<subseteq> S"
   13.39 +    using `open S` unfolding open_ereal_def by auto
   13.40 +  have "ereal -` uminus ` S = uminus ` (ereal -` S)"
   13.41    proof safe
   13.42 -    fix x y assume "extreal x = - y" "y \<in> S"
   13.43 -    then show "x \<in> uminus ` extreal -` S" by (cases y) auto
   13.44 +    fix x y assume "ereal x = - y" "y \<in> S"
   13.45 +    then show "x \<in> uminus ` ereal -` S" by (cases y) auto
   13.46    next
   13.47 -    fix x assume "extreal x \<in> S"
   13.48 -    then show "- x \<in> extreal -` uminus ` S"
   13.49 -      by (auto intro: image_eqI[of _ _ "extreal x"])
   13.50 +    fix x assume "ereal x \<in> S"
   13.51 +    then show "- x \<in> ereal -` uminus ` S"
   13.52 +      by (auto intro: image_eqI[of _ _ "ereal x"])
   13.53    qed
   13.54 -  then show "open (extreal -` uminus ` S)"
   13.55 +  then show "open (ereal -` uminus ` S)"
   13.56      using S by (auto intro: open_negations)
   13.57    { assume "\<infinity> \<in> uminus ` S"
   13.58 -    then have "-\<infinity> \<in> S" by (metis image_iff extreal_uminus_uminus)
   13.59 -    then have "uminus ` {..<extreal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
   13.60 -    then show "\<exists>x. {extreal x<..} \<subseteq> uminus ` S" using extreal_uminus_lessThan by auto }
   13.61 +    then have "-\<infinity> \<in> S" by (metis image_iff ereal_uminus_uminus)
   13.62 +    then have "uminus ` {..<ereal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
   13.63 +    then show "\<exists>x. {ereal x<..} \<subseteq> uminus ` S" using ereal_uminus_lessThan by auto }
   13.64    { assume "-\<infinity> \<in> uminus ` S"
   13.65 -    then have "\<infinity> : S" by (metis image_iff extreal_uminus_uminus)
   13.66 -    then have "uminus ` {extreal x<..} <= uminus ` S" using S by (intro image_mono) auto
   13.67 -    then show "\<exists>y. {..<extreal y} <= uminus ` S" using extreal_uminus_greaterThan by auto }
   13.68 +    then have "\<infinity> : S" by (metis image_iff ereal_uminus_uminus)
   13.69 +    then have "uminus ` {ereal x<..} <= uminus ` S" using S by (intro image_mono) auto
   13.70 +    then show "\<exists>y. {..<ereal y} <= uminus ` S" using ereal_uminus_greaterThan by auto }
   13.71  qed
   13.72  
   13.73 -lemma extreal_uminus_complement:
   13.74 -  fixes S :: "extreal set"
   13.75 +lemma ereal_uminus_complement:
   13.76 +  fixes S :: "ereal set"
   13.77    shows "uminus ` (- S) = - uminus ` S"
   13.78    by (auto intro!: bij_image_Compl_eq surjI[of _ uminus] simp: bij_betw_def)
   13.79  
   13.80 -lemma extreal_closed_uminus:
   13.81 -  fixes S :: "extreal set"
   13.82 +lemma ereal_closed_uminus:
   13.83 +  fixes S :: "ereal set"
   13.84    assumes "closed S"
   13.85    shows "closed (uminus ` S)"
   13.86  using assms unfolding closed_def
   13.87 -using extreal_open_uminus[of "- S"] extreal_uminus_complement by auto
   13.88 +using ereal_open_uminus[of "- S"] ereal_uminus_complement by auto
   13.89  
   13.90 -lemma not_open_extreal_singleton:
   13.91 -  "\<not> (open {a :: extreal})"
   13.92 +lemma not_open_ereal_singleton:
   13.93 +  "\<not> (open {a :: ereal})"
   13.94  proof(rule ccontr)
   13.95    assume "\<not> \<not> open {a}" hence a: "open {a}" by auto
   13.96    show False
   13.97    proof (cases a)
   13.98      case MInf
   13.99 -    then obtain y where "{..<extreal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
  13.100 -    hence "extreal(y - 1):{a}" apply (subst subsetD[of "{..<extreal y}"]) by auto
  13.101 +    then obtain y where "{..<ereal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
  13.102 +    hence "ereal(y - 1):{a}" apply (subst subsetD[of "{..<ereal y}"]) by auto
  13.103      then show False using `a=(-\<infinity>)` by auto
  13.104    next
  13.105      case PInf
  13.106 -    then obtain y where "{extreal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
  13.107 -    hence "extreal(y+1):{a}" apply (subst subsetD[of "{extreal y<..}"]) by auto
  13.108 +    then obtain y where "{ereal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
  13.109 +    hence "ereal(y+1):{a}" apply (subst subsetD[of "{ereal y<..}"]) by auto
  13.110      then show False using `a=\<infinity>` by auto
  13.111    next
  13.112      case (real r) then have fin: "\<bar>a\<bar> \<noteq> \<infinity>" by simp
  13.113 -    from extreal_open_cont_interval[OF a singletonI this] guess e . note e = this
  13.114 +    from ereal_open_cont_interval[OF a singletonI this] guess e . note e = this
  13.115      then obtain b where b_def: "a<b & b<a+e"
  13.116 -      using fin extreal_between extreal_dense[of a "a+e"] by auto
  13.117 -    then have "b: {a-e <..< a+e}" using fin extreal_between[of a e] e by auto
  13.118 +      using fin ereal_between ereal_dense[of a "a+e"] by auto
  13.119 +    then have "b: {a-e <..< a+e}" using fin ereal_between[of a e] e by auto
  13.120      then show False using b_def e by auto
  13.121    qed
  13.122  qed
  13.123  
  13.124 -lemma extreal_closed_contains_Inf:
  13.125 -  fixes S :: "extreal set"
  13.126 +lemma ereal_closed_contains_Inf:
  13.127 +  fixes S :: "ereal set"
  13.128    assumes "closed S" "S ~= {}"
  13.129    shows "Inf S : S"
  13.130  proof(rule ccontr)
  13.131 @@ -96,8 +96,8 @@
  13.132    show False
  13.133    proof (cases "Inf S")
  13.134      case MInf hence "(-\<infinity>) : - S" using a by auto
  13.135 -    then obtain y where "{..<extreal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
  13.136 -    hence "extreal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff
  13.137 +    then obtain y where "{..<ereal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
  13.138 +    hence "ereal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff
  13.139        complete_lattice_class.Inf_greatest double_complement set_rev_mp)
  13.140      then show False using MInf by auto
  13.141    next
  13.142 @@ -105,9 +105,9 @@
  13.143      then show False by (metis `Inf S ~: S` insert_code mem_def PInf)
  13.144    next
  13.145      case (real r) then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp
  13.146 -    from extreal_open_cont_interval[OF a this] guess e . note e = this
  13.147 +    from ereal_open_cont_interval[OF a this] guess e . note e = this
  13.148      { fix x assume "x:S" hence "x>=Inf S" by (rule complete_lattice_class.Inf_lower)
  13.149 -      hence *: "x>Inf S-e" using e by (metis fin extreal_between(1) order_less_le_trans)
  13.150 +      hence *: "x>Inf S-e" using e by (metis fin ereal_between(1) order_less_le_trans)
  13.151        { assume "x<Inf S+e" hence "x:{Inf S-e <..< Inf S+e}" using * by auto
  13.152          hence False using e `x:S` by auto
  13.153        } hence "x>=Inf S+e" by (metis linorder_le_less_linear)
  13.154 @@ -116,115 +116,117 @@
  13.155    qed
  13.156  qed
  13.157  
  13.158 -lemma extreal_closed_contains_Sup:
  13.159 -  fixes S :: "extreal set"
  13.160 +lemma ereal_closed_contains_Sup:
  13.161 +  fixes S :: "ereal set"
  13.162    assumes "closed S" "S ~= {}"
  13.163    shows "Sup S : S"
  13.164  proof-
  13.165 -  have "closed (uminus ` S)" by (metis assms(1) extreal_closed_uminus)
  13.166 -  hence "Inf (uminus ` S) : uminus ` S" using assms extreal_closed_contains_Inf[of "uminus ` S"] by auto
  13.167 -  hence "- Sup S : uminus ` S" using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image)
  13.168 -  thus ?thesis by (metis imageI extreal_uminus_uminus extreal_minus_minus_image)
  13.169 +  have "closed (uminus ` S)" by (metis assms(1) ereal_closed_uminus)
  13.170 +  hence "Inf (uminus ` S) : uminus ` S" using assms ereal_closed_contains_Inf[of "uminus ` S"] by auto
  13.171 +  hence "- Sup S : uminus ` S" using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image)
  13.172 +  thus ?thesis by (metis imageI ereal_uminus_uminus ereal_minus_minus_image)
  13.173  qed
  13.174  
  13.175 -lemma extreal_open_closed_aux:
  13.176 -  fixes S :: "extreal set"
  13.177 +lemma ereal_open_closed_aux:
  13.178 +  fixes S :: "ereal set"
  13.179    assumes "open S" "closed S"
  13.180    assumes S: "(-\<infinity>) ~: S"
  13.181    shows "S = {}"
  13.182  proof(rule ccontr)
  13.183    assume "S ~= {}"
  13.184 -  hence *: "(Inf S):S" by (metis assms(2) extreal_closed_contains_Inf)
  13.185 +  hence *: "(Inf S):S" by (metis assms(2) ereal_closed_contains_Inf)
  13.186    { assume "Inf S=(-\<infinity>)" hence False using * assms(3) by auto }
  13.187    moreover
  13.188    { assume "Inf S=\<infinity>" hence "S={\<infinity>}" by (metis Inf_eq_PInfty `S ~= {}`)
  13.189 -    hence False by (metis assms(1) not_open_extreal_singleton) }
  13.190 +    hence False by (metis assms(1) not_open_ereal_singleton) }
  13.191    moreover
  13.192    { assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"
  13.193 -    from extreal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
  13.194 +    from ereal_open_cont_interval[OF assms(1) * fin] guess e . note e = this
  13.195      then obtain b where b_def: "Inf S-e<b & b<Inf S"
  13.196 -      using fin extreal_between[of "Inf S" e] extreal_dense[of "Inf S-e"] by auto
  13.197 -    hence "b: {Inf S-e <..< Inf S+e}" using e fin extreal_between[of "Inf S" e] by auto
  13.198 +      using fin ereal_between[of "Inf S" e] ereal_dense[of "Inf S-e"] by auto
  13.199 +    hence "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e] by auto
  13.200      hence "b:S" using e by auto
  13.201      hence False using b_def by (metis complete_lattice_class.Inf_lower leD)
  13.202    } ultimately show False by auto
  13.203  qed
  13.204  
  13.205 -lemma extreal_open_closed:
  13.206 -  fixes S :: "extreal set"
  13.207 +lemma ereal_open_closed:
  13.208 +  fixes S :: "ereal set"
  13.209    shows "(open S & closed S) <-> (S = {} | S = UNIV)"
  13.210  proof-
  13.211  { assume lhs: "open S & closed S"
  13.212 -  { assume "(-\<infinity>) ~: S" hence "S={}" using lhs extreal_open_closed_aux by auto }
  13.213 +  { assume "(-\<infinity>) ~: S" hence "S={}" using lhs ereal_open_closed_aux by auto }
  13.214    moreover
  13.215 -  { assume "(-\<infinity>) : S" hence "(- S)={}" using lhs extreal_open_closed_aux[of "-S"] by auto }
  13.216 +  { assume "(-\<infinity>) : S" hence "(- S)={}" using lhs ereal_open_closed_aux[of "-S"] by auto }
  13.217    ultimately have "S = {} | S = UNIV" by auto
  13.218  } thus ?thesis by auto
  13.219  qed
  13.220  
  13.221 -lemma extreal_open_affinity_pos:
  13.222 +lemma ereal_open_affinity_pos:
  13.223 +  fixes S :: "ereal set"
  13.224    assumes "open S" and m: "m \<noteq> \<infinity>" "0 < m" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
  13.225    shows "open ((\<lambda>x. m * x + t) ` S)"
  13.226  proof -
  13.227 -  obtain r where r[simp]: "m = extreal r" using m by (cases m) auto
  13.228 -  obtain p where p[simp]: "t = extreal p" using t by auto
  13.229 +  obtain r where r[simp]: "m = ereal r" using m by (cases m) auto
  13.230 +  obtain p where p[simp]: "t = ereal p" using t by auto
  13.231    have "r \<noteq> 0" "0 < r" and m': "m \<noteq> \<infinity>" "m \<noteq> -\<infinity>" "m \<noteq> 0" using m by auto
  13.232 -  from `open S`[THEN extreal_openE] guess l u . note T = this
  13.233 +  from `open S`[THEN ereal_openE] guess l u . note T = this
  13.234    let ?f = "(\<lambda>x. m * x + t)"
  13.235 -  show ?thesis unfolding open_extreal_def
  13.236 +  show ?thesis unfolding open_ereal_def
  13.237    proof (intro conjI impI exI subsetI)
  13.238 -    have "extreal -` ?f ` S = (\<lambda>x. r * x + p) ` (extreal -` S)"
  13.239 +    have "ereal -` ?f ` S = (\<lambda>x. r * x + p) ` (ereal -` S)"
  13.240      proof safe
  13.241 -      fix x y assume "extreal y = m * x + t" "x \<in> S"
  13.242 -      then show "y \<in> (\<lambda>x. r * x + p) ` extreal -` S"
  13.243 +      fix x y assume "ereal y = m * x + t" "x \<in> S"
  13.244 +      then show "y \<in> (\<lambda>x. r * x + p) ` ereal -` S"
  13.245          using `r \<noteq> 0` by (cases x) (auto intro!: image_eqI[of _ _ "real x"] split: split_if_asm)
  13.246      qed force
  13.247 -    then show "open (extreal -` ?f ` S)"
  13.248 +    then show "open (ereal -` ?f ` S)"
  13.249        using open_affinity[OF T(1) `r \<noteq> 0`] by (auto simp: ac_simps)
  13.250    next
  13.251      assume "\<infinity> \<in> ?f`S" with `0 < r` have "\<infinity> \<in> S" by auto
  13.252 -    fix x assume "x \<in> {extreal (r * l + p)<..}"
  13.253 -    then have [simp]: "extreal (r * l + p) < x" by auto
  13.254 +    fix x assume "x \<in> {ereal (r * l + p)<..}"
  13.255 +    then have [simp]: "ereal (r * l + p) < x" by auto
  13.256      show "x \<in> ?f`S"
  13.257      proof (rule image_eqI)
  13.258        show "x = m * ((x - t) / m) + t"
  13.259 -        using m t by (cases rule: extreal3_cases[of m x t]) auto
  13.260 -      have "extreal l < (x - t)/m"
  13.261 -        using m t by (simp add: extreal_less_divide_pos extreal_less_minus)
  13.262 +        using m t by (cases rule: ereal3_cases[of m x t]) auto
  13.263 +      have "ereal l < (x - t)/m"
  13.264 +        using m t by (simp add: ereal_less_divide_pos ereal_less_minus)
  13.265        then show "(x - t)/m \<in> S" using T(2)[OF `\<infinity> \<in> S`] by auto
  13.266      qed
  13.267    next
  13.268      assume "-\<infinity> \<in> ?f`S" with `0 < r` have "-\<infinity> \<in> S" by auto
  13.269 -    fix x assume "x \<in> {..<extreal (r * u + p)}"
  13.270 -    then have [simp]: "x < extreal (r * u + p)" by auto
  13.271 +    fix x assume "x \<in> {..<ereal (r * u + p)}"
  13.272 +    then have [simp]: "x < ereal (r * u + p)" by auto
  13.273      show "x \<in> ?f`S"
  13.274      proof (rule image_eqI)
  13.275        show "x = m * ((x - t) / m) + t"
  13.276 -        using m t by (cases rule: extreal3_cases[of m x t]) auto
  13.277 -      have "(x - t)/m < extreal u"
  13.278 -        using m t by (simp add: extreal_divide_less_pos extreal_minus_less)
  13.279 +        using m t by (cases rule: ereal3_cases[of m x t]) auto
  13.280 +      have "(x - t)/m < ereal u"
  13.281 +        using m t by (simp add: ereal_divide_less_pos ereal_minus_less)
  13.282        then show "(x - t)/m \<in> S" using T(3)[OF `-\<infinity> \<in> S`] by auto
  13.283      qed
  13.284    qed
  13.285  qed
  13.286  
  13.287 -lemma extreal_open_affinity:
  13.288 +lemma ereal_open_affinity:
  13.289 +  fixes S :: "ereal set"
  13.290    assumes "open S" and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" and t: "\<bar>t\<bar> \<noteq> \<infinity>"
  13.291    shows "open ((\<lambda>x. m * x + t) ` S)"
  13.292  proof cases
  13.293    assume "0 < m" then show ?thesis
  13.294 -    using extreal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto
  13.295 +    using ereal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto
  13.296  next
  13.297    assume "\<not> 0 < m" then
  13.298    have "0 < -m" using `m \<noteq> 0` by (cases m) auto
  13.299    then have m: "-m \<noteq> \<infinity>" "0 < -m" using `\<bar>m\<bar> \<noteq> \<infinity>`
  13.300 -    by (auto simp: extreal_uminus_eq_reorder)
  13.301 -  from extreal_open_affinity_pos[OF extreal_open_uminus[OF `open S`] m t]
  13.302 +    by (auto simp: ereal_uminus_eq_reorder)
  13.303 +  from ereal_open_affinity_pos[OF ereal_open_uminus[OF `open S`] m t]
  13.304    show ?thesis unfolding image_image by simp
  13.305  qed
  13.306  
  13.307 -lemma extreal_lim_mult:
  13.308 -  fixes X :: "'a \<Rightarrow> extreal"
  13.309 +lemma ereal_lim_mult:
  13.310 +  fixes X :: "'a \<Rightarrow> ereal"
  13.311    assumes lim: "(X ---> L) net" and a: "\<bar>a\<bar> \<noteq> \<infinity>"
  13.312    shows "((\<lambda>i. a * X i) ---> a * L) net"
  13.313  proof cases
  13.314 @@ -233,73 +235,73 @@
  13.315    proof (rule topological_tendstoI)
  13.316      fix S assume "open S" "a * L \<in> S"
  13.317      have "a * L / a = L"
  13.318 -      using `a \<noteq> 0` a by (cases rule: extreal2_cases[of a L]) auto
  13.319 +      using `a \<noteq> 0` a by (cases rule: ereal2_cases[of a L]) auto
  13.320      then have L: "L \<in> ((\<lambda>x. x / a) ` S)"
  13.321        using `a * L \<in> S` by (force simp: image_iff)
  13.322      moreover have "open ((\<lambda>x. x / a) ` S)"
  13.323 -      using extreal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a
  13.324 -      by (auto simp: extreal_divide_eq extreal_inverse_eq_0 divide_extreal_def ac_simps)
  13.325 +      using ereal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a
  13.326 +      by (auto simp: ereal_divide_eq ereal_inverse_eq_0 divide_ereal_def ac_simps)
  13.327      note * = lim[THEN topological_tendstoD, OF this L]
  13.328      { fix x from a `a \<noteq> 0` have "a * (x / a) = x"
  13.329 -        by (cases rule: extreal2_cases[of a x]) auto }
  13.330 +        by (cases rule: ereal2_cases[of a x]) auto }
  13.331      note this[simp]
  13.332      show "eventually (\<lambda>x. a * X x \<in> S) net"
  13.333        by (rule eventually_mono[OF _ *]) auto
  13.334    qed
  13.335  qed auto
  13.336  
  13.337 -lemma extreal_lim_uminus:
  13.338 -  fixes X :: "'a \<Rightarrow> extreal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
  13.339 -  using extreal_lim_mult[of X L net "extreal (-1)"]
  13.340 -        extreal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "extreal (-1)"]
  13.341 +lemma ereal_lim_uminus:
  13.342 +  fixes X :: "'a \<Rightarrow> ereal" shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net"
  13.343 +  using ereal_lim_mult[of X L net "ereal (-1)"]
  13.344 +        ereal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "ereal (-1)"]
  13.345    by (auto simp add: algebra_simps)
  13.346  
  13.347 -lemma Lim_bounded2_extreal:
  13.348 -  assumes lim:"f ----> (l :: extreal)"
  13.349 +lemma Lim_bounded2_ereal:
  13.350 +  assumes lim:"f ----> (l :: ereal)"
  13.351    and ge: "ALL n>=N. f n >= C"
  13.352    shows "l>=C"
  13.353  proof-
  13.354  def g == "(%i. -(f i))"
  13.355 -{ fix n assume "n>=N" hence "g n <= -C" using assms extreal_minus_le_minus g_def by auto }
  13.356 +{ fix n assume "n>=N" hence "g n <= -C" using assms ereal_minus_le_minus g_def by auto }
  13.357  hence "ALL n>=N. g n <= -C" by auto
  13.358 -moreover have limg: "g ----> (-l)" using g_def extreal_lim_uminus lim by auto
  13.359 -ultimately have "-l <= -C" using Lim_bounded_extreal[of g "-l" _ "-C"] by auto
  13.360 -from this show ?thesis using extreal_minus_le_minus by auto
  13.361 +moreover have limg: "g ----> (-l)" using g_def ereal_lim_uminus lim by auto
  13.362 +ultimately have "-l <= -C" using Lim_bounded_ereal[of g "-l" _ "-C"] by auto
  13.363 +from this show ?thesis using ereal_minus_le_minus by auto
  13.364  qed
  13.365  
  13.366  
  13.367 -lemma extreal_open_atLeast: "open {x..} \<longleftrightarrow> x = -\<infinity>"
  13.368 +lemma ereal_open_atLeast: fixes x :: ereal shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
  13.369  proof
  13.370    assume "x = -\<infinity>" then have "{x..} = UNIV" by auto
  13.371    then show "open {x..}" by auto
  13.372  next
  13.373    assume "open {x..}"
  13.374    then have "open {x..} \<and> closed {x..}" by auto
  13.375 -  then have "{x..} = UNIV" unfolding extreal_open_closed by auto
  13.376 -  then show "x = -\<infinity>" by (simp add: bot_extreal_def atLeast_eq_UNIV_iff)
  13.377 +  then have "{x..} = UNIV" unfolding ereal_open_closed by auto
  13.378 +  then show "x = -\<infinity>" by (simp add: bot_ereal_def atLeast_eq_UNIV_iff)
  13.379  qed
  13.380  
  13.381 -lemma extreal_open_mono_set:
  13.382 -  fixes S :: "extreal set"
  13.383 +lemma ereal_open_mono_set:
  13.384 +  fixes S :: "ereal set"
  13.385    defines "a \<equiv> Inf S"
  13.386    shows "(open S \<and> mono S) \<longleftrightarrow> (S = UNIV \<or> S = {a <..})"
  13.387 -  by (metis Inf_UNIV a_def atLeast_eq_UNIV_iff extreal_open_atLeast
  13.388 -            extreal_open_closed mono_set_iff open_extreal_greaterThan)
  13.389 +  by (metis Inf_UNIV a_def atLeast_eq_UNIV_iff ereal_open_atLeast
  13.390 +            ereal_open_closed mono_set_iff open_ereal_greaterThan)
  13.391  
  13.392 -lemma extreal_closed_mono_set:
  13.393 -  fixes S :: "extreal set"
  13.394 +lemma ereal_closed_mono_set:
  13.395 +  fixes S :: "ereal set"
  13.396    shows "(closed S \<and> mono S) \<longleftrightarrow> (S = {} \<or> S = {Inf S ..})"
  13.397 -  by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_extreal_atLeast
  13.398 -            extreal_open_closed mono_empty mono_set_iff open_extreal_greaterThan)
  13.399 +  by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast
  13.400 +            ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan)
  13.401  
  13.402 -lemma extreal_Liminf_Sup_monoset:
  13.403 -  fixes f :: "'a => extreal"
  13.404 +lemma ereal_Liminf_Sup_monoset:
  13.405 +  fixes f :: "'a => ereal"
  13.406    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}"
  13.407    unfolding Liminf_Sup
  13.408  proof (intro arg_cong[where f="\<lambda>P. Sup (Collect P)"] ext iffI allI impI)
  13.409    fix l S assume ev: "\<forall>y<l. eventually (\<lambda>x. y < f x) net" and "open S" "mono S" "l \<in> S"
  13.410    then have "S = UNIV \<or> S = {Inf S <..}"
  13.411 -    using extreal_open_mono_set[of S] by auto
  13.412 +    using ereal_open_mono_set[of S] by auto
  13.413    then show "eventually (\<lambda>x. f x \<in> S) net"
  13.414    proof
  13.415      assume S: "S = {Inf S<..}"
  13.416 @@ -314,15 +316,15 @@
  13.417    then show "eventually (\<lambda>x. y < f x) net" by auto
  13.418  qed
  13.419  
  13.420 -lemma extreal_Limsup_Inf_monoset:
  13.421 -  fixes f :: "'a => extreal"
  13.422 +lemma ereal_Limsup_Inf_monoset:
  13.423 +  fixes f :: "'a => ereal"
  13.424    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}"
  13.425    unfolding Limsup_Inf
  13.426  proof (intro arg_cong[where f="\<lambda>P. Inf (Collect P)"] ext iffI allI impI)
  13.427    fix l S assume ev: "\<forall>y>l. eventually (\<lambda>x. f x < y) net" and "open S" "mono (uminus`S)" "l \<in> S"
  13.428 -  then have "open (uminus`S) \<and> mono (uminus`S)" by (simp add: extreal_open_uminus)
  13.429 +  then have "open (uminus`S) \<and> mono (uminus`S)" by (simp add: ereal_open_uminus)
  13.430    then have "S = UNIV \<or> S = {..< Sup S}"
  13.431 -    unfolding extreal_open_mono_set extreal_Inf_uminus_image_eq extreal_image_uminus_shift by simp
  13.432 +    unfolding ereal_open_mono_set ereal_Inf_uminus_image_eq ereal_image_uminus_shift by simp
  13.433    then show "eventually (\<lambda>x. f x \<in> S) net"
  13.434    proof
  13.435      assume S: "S = {..< Sup S}"
  13.436 @@ -338,70 +340,70 @@
  13.437  qed
  13.438  
  13.439  
  13.440 -lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::extreal set)"
  13.441 -  using extreal_open_uminus[of S] extreal_open_uminus[of "uminus`S"] by auto
  13.442 +lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::ereal set)"
  13.443 +  using ereal_open_uminus[of S] ereal_open_uminus[of "uminus`S"] by auto
  13.444  
  13.445 -lemma extreal_Limsup_uminus:
  13.446 -  fixes f :: "'a => extreal"
  13.447 +lemma ereal_Limsup_uminus:
  13.448 +  fixes f :: "'a => ereal"
  13.449    shows "Limsup net (\<lambda>x. - (f x)) = -(Liminf net f)"
  13.450  proof -
  13.451 -  { fix P l have "(\<exists>x. (l::extreal) = -x \<and> P x) \<longleftrightarrow> P (-l)" by (auto intro!: exI[of _ "-l"]) }
  13.452 +  { fix P l have "(\<exists>x. (l::ereal) = -x \<and> P x) \<longleftrightarrow> P (-l)" by (auto intro!: exI[of _ "-l"]) }
  13.453    note Ex_cancel = this
  13.454 -  { fix P :: "extreal set \<Rightarrow> bool" have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))"
  13.455 +  { fix P :: "ereal set \<Rightarrow> bool" have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))"
  13.456        apply auto by (erule_tac x="uminus`S" in allE) (auto simp: image_image) }
  13.457    note add_uminus_image = this
  13.458 -  { fix x S have "(x::extreal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" by (auto intro!: image_eqI[of _ _ "-x"]) }
  13.459 +  { fix x S have "(x::ereal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" by (auto intro!: image_eqI[of _ _ "-x"]) }
  13.460    note remove_uminus_image = this
  13.461    show ?thesis
  13.462 -    unfolding extreal_Limsup_Inf_monoset extreal_Liminf_Sup_monoset
  13.463 -    unfolding extreal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel
  13.464 +    unfolding ereal_Limsup_Inf_monoset ereal_Liminf_Sup_monoset
  13.465 +    unfolding ereal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel
  13.466      by (subst add_uminus_image) (simp add: open_uminus_iff remove_uminus_image)
  13.467  qed
  13.468  
  13.469 -lemma extreal_Liminf_uminus:
  13.470 -  fixes f :: "'a => extreal"
  13.471 +lemma ereal_Liminf_uminus:
  13.472 +  fixes f :: "'a => ereal"
  13.473    shows "Liminf net (\<lambda>x. - (f x)) = -(Limsup net f)"
  13.474 -  using extreal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto
  13.475 +  using ereal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto
  13.476  
  13.477 -lemma extreal_Lim_uminus:
  13.478 -  fixes f :: "'a \<Rightarrow> extreal" shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net"
  13.479 +lemma ereal_Lim_uminus:
  13.480 +  fixes f :: "'a \<Rightarrow> ereal" shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net"
  13.481    using
  13.482 -    extreal_lim_mult[of f f0 net "- 1"]
  13.483 -    extreal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"]
  13.484 -  by (auto simp: extreal_uminus_reorder)
  13.485 +    ereal_lim_mult[of f f0 net "- 1"]
  13.486 +    ereal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"]
  13.487 +  by (auto simp: ereal_uminus_reorder)
  13.488  
  13.489  lemma lim_imp_Limsup:
  13.490 -  fixes f :: "'a => extreal"
  13.491 +  fixes f :: "'a => ereal"
  13.492    assumes "\<not> trivial_limit net"
  13.493    assumes lim: "(f ---> f0) net"
  13.494    shows "Limsup net f = f0"
  13.495 -  using extreal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"]
  13.496 -     extreal_Liminf_uminus[of net f] assms by simp
  13.497 +  using ereal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"]
  13.498 +     ereal_Liminf_uminus[of net f] assms by simp
  13.499  
  13.500  lemma Liminf_PInfty:
  13.501 -  fixes f :: "'a \<Rightarrow> extreal"
  13.502 +  fixes f :: "'a \<Rightarrow> ereal"
  13.503    assumes "\<not> trivial_limit net"
  13.504    shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>"
  13.505  proof (intro lim_imp_Liminf iffI assms)
  13.506    assume rhs: "Liminf net f = \<infinity>"
  13.507 -  { fix S assume "open S & \<infinity> : S"
  13.508 -    then obtain m where "{extreal m<..} <= S" using open_PInfty2 by auto
  13.509 -    moreover have "eventually (\<lambda>x. f x \<in> {extreal m<..}) net"
  13.510 -      using rhs unfolding Liminf_Sup top_extreal_def[symmetric] Sup_eq_top_iff
  13.511 -      by (auto elim!: allE[where x="extreal m"] simp: top_extreal_def)
  13.512 +  { fix S :: "ereal set" assume "open S & \<infinity> : S"
  13.513 +    then obtain m where "{ereal m<..} <= S" using open_PInfty2 by auto
  13.514 +    moreover have "eventually (\<lambda>x. f x \<in> {ereal m<..}) net"
  13.515 +      using rhs unfolding Liminf_Sup top_ereal_def[symmetric] Sup_eq_top_iff
  13.516 +      by (auto elim!: allE[where x="ereal m"] simp: top_ereal_def)
  13.517      ultimately have "eventually (%x. f x : S) net" apply (subst eventually_mono) by auto
  13.518    } then show "(f ---> \<infinity>) net" unfolding tendsto_def by auto
  13.519  qed
  13.520  
  13.521  lemma Limsup_MInfty:
  13.522 -  fixes f :: "'a \<Rightarrow> extreal"
  13.523 +  fixes f :: "'a \<Rightarrow> ereal"
  13.524    assumes "\<not> trivial_limit net"
  13.525    shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>"
  13.526 -  using assms extreal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"]
  13.527 -        extreal_Liminf_uminus[of _ f] by (auto simp: extreal_uminus_eq_reorder)
  13.528 +  using assms ereal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"]
  13.529 +        ereal_Liminf_uminus[of _ f] by (auto simp: ereal_uminus_eq_reorder)
  13.530  
  13.531 -lemma extreal_Liminf_eq_Limsup:
  13.532 -  fixes f :: "'a \<Rightarrow> extreal"
  13.533 +lemma ereal_Liminf_eq_Limsup:
  13.534 +  fixes f :: "'a \<Rightarrow> ereal"
  13.535    assumes ntriv: "\<not> trivial_limit net"
  13.536    assumes lim: "Liminf net f = f0" "Limsup net f = f0"
  13.537    shows "(f ---> f0) net"
  13.538 @@ -415,7 +417,7 @@
  13.539    proof (rule topological_tendstoI)
  13.540      fix S assume "open S""f0 \<in> S"
  13.541      then obtain a b where "a < Liminf net f" "Limsup net f < b" "{a<..<b} \<subseteq> S"
  13.542 -      using extreal_open_cont_interval2[of S f0] real lim by auto
  13.543 +      using ereal_open_cont_interval2[of S f0] real lim by auto
  13.544      then have "eventually (\<lambda>x. f x \<in> {a<..<b}) net"
  13.545        unfolding Liminf_Sup Limsup_Inf less_Sup_iff Inf_less_iff
  13.546        by (auto intro!: eventually_conj simp add: greaterThanLessThan_iff)
  13.547 @@ -424,62 +426,62 @@
  13.548    qed
  13.549  qed
  13.550  
  13.551 -lemma extreal_Liminf_eq_Limsup_iff:
  13.552 -  fixes f :: "'a \<Rightarrow> extreal"
  13.553 +lemma ereal_Liminf_eq_Limsup_iff:
  13.554 +  fixes f :: "'a \<Rightarrow> ereal"
  13.555    assumes "\<not> trivial_limit net"
  13.556    shows "(f ---> f0) net \<longleftrightarrow> Liminf net f = f0 \<and> Limsup net f = f0"
  13.557 -  by (metis assms extreal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup)
  13.558 +  by (metis assms ereal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup)
  13.559  
  13.560  lemma limsup_INFI_SUPR:
  13.561 -  fixes f :: "nat \<Rightarrow> extreal"
  13.562 +  fixes f :: "nat \<Rightarrow> ereal"
  13.563    shows "limsup f = (INF n. SUP m:{n..}. f m)"
  13.564 -  using extreal_Limsup_uminus[of sequentially "\<lambda>x. - f x"]
  13.565 -  by (simp add: liminf_SUPR_INFI extreal_INFI_uminus extreal_SUPR_uminus)
  13.566 +  using ereal_Limsup_uminus[of sequentially "\<lambda>x. - f x"]
  13.567 +  by (simp add: liminf_SUPR_INFI ereal_INFI_uminus ereal_SUPR_uminus)
  13.568  
  13.569  lemma liminf_PInfty:
  13.570 -  fixes X :: "nat => extreal"
  13.571 +  fixes X :: "nat => ereal"
  13.572    shows "X ----> \<infinity> <-> liminf X = \<infinity>"
  13.573  by (metis Liminf_PInfty trivial_limit_sequentially)
  13.574  
  13.575  lemma limsup_MInfty:
  13.576 -  fixes X :: "nat => extreal"
  13.577 +  fixes X :: "nat => ereal"
  13.578    shows "X ----> (-\<infinity>) <-> limsup X = (-\<infinity>)"
  13.579  by (metis Limsup_MInfty trivial_limit_sequentially)
  13.580  
  13.581 -lemma extreal_lim_mono:
  13.582 -  fixes X Y :: "nat => extreal"
  13.583 +lemma ereal_lim_mono:
  13.584 +  fixes X Y :: "nat => ereal"
  13.585    assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n"
  13.586    assumes "X ----> x" "Y ----> y"
  13.587    shows "x <= y"
  13.588 -  by (metis extreal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
  13.589 +  by (metis ereal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono)
  13.590  
  13.591 -lemma incseq_le_extreal:
  13.592 -  fixes X :: "nat \<Rightarrow> extreal"
  13.593 +lemma incseq_le_ereal:
  13.594 +  fixes X :: "nat \<Rightarrow> ereal"
  13.595    assumes inc: "incseq X" and lim: "X ----> L"
  13.596    shows "X N \<le> L"
  13.597    using inc
  13.598 -  by (intro extreal_lim_mono[of N, OF _ Lim_const lim]) (simp add: incseq_def)
  13.599 +  by (intro ereal_lim_mono[of N, OF _ Lim_const lim]) (simp add: incseq_def)
  13.600  
  13.601 -lemma decseq_ge_extreal: assumes dec: "decseq X"
  13.602 -  and lim: "X ----> (L::extreal)" shows "X N >= L"
  13.603 +lemma decseq_ge_ereal: assumes dec: "decseq X"
  13.604 +  and lim: "X ----> (L::ereal)" shows "X N >= L"
  13.605    using dec
  13.606 -  by (intro extreal_lim_mono[of N, OF _ lim Lim_const]) (simp add: decseq_def)
  13.607 +  by (intro ereal_lim_mono[of N, OF _ lim Lim_const]) (simp add: decseq_def)
  13.608  
  13.609  lemma liminf_bounded_open:
  13.610 -  fixes x :: "nat \<Rightarrow> extreal"
  13.611 +  fixes x :: "nat \<Rightarrow> ereal"
  13.612    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))" 
  13.613    (is "_ \<longleftrightarrow> ?P x0")
  13.614  proof
  13.615    assume "?P x0" then show "x0 \<le> liminf x"
  13.616 -    unfolding extreal_Liminf_Sup_monoset eventually_sequentially
  13.617 +    unfolding ereal_Liminf_Sup_monoset eventually_sequentially
  13.618      by (intro complete_lattice_class.Sup_upper) auto
  13.619  next
  13.620    assume "x0 \<le> liminf x"
  13.621 -  { fix S :: "extreal set" assume om: "open S & mono S & x0:S"
  13.622 +  { fix S :: "ereal set" assume om: "open S & mono S & x0:S"
  13.623      { assume "S = UNIV" hence "EX N. (ALL n>=N. x n : S)" by auto }
  13.624      moreover
  13.625      { assume "~(S=UNIV)"
  13.626 -      then obtain B where B_def: "S = {B<..}" using om extreal_open_mono_set by auto
  13.627 +      then obtain B where B_def: "S = {B<..}" using om ereal_open_mono_set by auto
  13.628        hence "B<x0" using om by auto
  13.629        hence "EX N. ALL n>=N. x n : S" unfolding B_def using `x0 \<le> liminf x` liminf_bounded_iff by auto
  13.630      } ultimately have "EX N. (ALL n>=N. x n : S)" by auto
  13.631 @@ -487,15 +489,15 @@
  13.632  qed
  13.633  
  13.634  lemma limsup_subseq_mono:
  13.635 -  fixes X :: "nat \<Rightarrow> extreal"
  13.636 +  fixes X :: "nat \<Rightarrow> ereal"
  13.637    assumes "subseq r"
  13.638    shows "limsup (X \<circ> r) \<le> limsup X"
  13.639  proof-
  13.640    have "(\<lambda>n. - X n) \<circ> r = (\<lambda>n. - (X \<circ> r) n)" by (simp add: fun_eq_iff)
  13.641    then have "- limsup X \<le> - limsup (X \<circ> r)"
  13.642       using liminf_subseq_mono[of r "(%n. - X n)"]
  13.643 -       extreal_Liminf_uminus[of sequentially X]
  13.644 -       extreal_Liminf_uminus[of sequentially "X o r"] assms by auto
  13.645 +       ereal_Liminf_uminus[of sequentially X]
  13.646 +       ereal_Liminf_uminus[of sequentially "X o r"] assms by auto
  13.647    then show ?thesis by auto
  13.648  qed
  13.649  
  13.650 @@ -514,8 +516,8 @@
  13.651  from this show ?thesis apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
  13.652     using assms by auto
  13.653  qed
  13.654 -lemma lim_extreal_increasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
  13.655 -  obtains l where "f ----> (l::extreal)"
  13.656 +lemma lim_ereal_increasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m"
  13.657 +  obtains l where "f ----> (l::ereal)"
  13.658  proof(cases "f = (\<lambda>x. - \<infinity>)")
  13.659    case True then show thesis using Lim_const[of "- \<infinity>" sequentially] by (intro that[of "-\<infinity>"]) auto
  13.660  next
  13.661 @@ -527,18 +529,18 @@
  13.662    hence incy: "!!n m. n>=m ==> Y n >= Y m" using assms by auto
  13.663    from minf have minfy: "ALL n. Y n ~= (-\<infinity>)" using Y_def by auto
  13.664    show thesis
  13.665 -  proof(cases "EX B. ALL n. f n < extreal B")
  13.666 +  proof(cases "EX B. ALL n. f n < ereal B")
  13.667      case False thus thesis apply- apply(rule that[of \<infinity>]) unfolding Lim_PInfty not_ex not_all
  13.668      apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
  13.669      apply(rule order_trans[OF _ assms[rule_format]]) by auto
  13.670    next case True then guess B ..
  13.671 -    hence "ALL n. Y n < extreal B" using Y_def by auto note B = this[rule_format]
  13.672 +    hence "ALL n. Y n < ereal B" using Y_def by auto note B = this[rule_format]
  13.673      { fix n have "Y n < \<infinity>" using B[of n] apply (subst less_le_trans) by auto
  13.674        hence "Y n ~= \<infinity> & Y n ~= (-\<infinity>)" using minfy by auto
  13.675      } hence *: "ALL n. \<bar>Y n\<bar> \<noteq> \<infinity>" by auto
  13.676      { fix n have "real (Y n) < B" proof- case goal1 thus ?case
  13.677 -        using B[of n] apply-apply(subst(asm) extreal_real'[THEN sym]) defer defer
  13.678 -        unfolding extreal_less using * by auto
  13.679 +        using B[of n] apply-apply(subst(asm) ereal_real'[THEN sym]) defer defer
  13.680 +        unfolding ereal_less using * by auto
  13.681        qed
  13.682      }
  13.683      hence B': "ALL n. (real (Y n) <= B)" using less_imp_le by auto
  13.684 @@ -546,29 +548,29 @@
  13.685        apply(rule bounded_increasing_convergent2)
  13.686      proof safe show "!!n. real (Y n) <= B" using B' by auto
  13.687        fix n m::nat assume "n<=m"
  13.688 -      hence "extreal (real (Y n)) <= extreal (real (Y m))"
  13.689 -        using incy[rule_format,of n m] apply(subst extreal_real)+
  13.690 +      hence "ereal (real (Y n)) <= ereal (real (Y m))"
  13.691 +        using incy[rule_format,of n m] apply(subst ereal_real)+
  13.692          using *[rule_format, of n] *[rule_format, of m] by auto
  13.693        thus "real (Y n) <= real (Y m)" by auto
  13.694      qed then guess l .. note l=this
  13.695 -    have "Y ----> extreal l" using l apply-apply(subst(asm) lim_extreal[THEN sym])
  13.696 -    unfolding extreal_real using * by auto
  13.697 -    thus thesis apply-apply(rule that[of "extreal l"])
  13.698 +    have "Y ----> ereal l" using l apply-apply(subst(asm) lim_ereal[THEN sym])
  13.699 +    unfolding ereal_real using * by auto
  13.700 +    thus thesis apply-apply(rule that[of "ereal l"])
  13.701         apply (subst tail_same_limit[of Y _ N]) using Y_def by auto
  13.702    qed
  13.703  qed
  13.704  
  13.705 -lemma lim_extreal_decreasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
  13.706 -  obtains l where "f ----> (l::extreal)"
  13.707 +lemma lim_ereal_decreasing: assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m"
  13.708 +  obtains l where "f ----> (l::ereal)"
  13.709  proof -
  13.710 -  from lim_extreal_increasing[of "\<lambda>x. - f x"] assms
  13.711 +  from lim_ereal_increasing[of "\<lambda>x. - f x"] assms
  13.712    obtain l where "(\<lambda>x. - f x) ----> l" by auto
  13.713 -  from extreal_lim_mult[OF this, of "- 1"] show thesis
  13.714 -    by (intro that[of "-l"]) (simp add: extreal_uminus_eq_reorder)
  13.715 +  from ereal_lim_mult[OF this, of "- 1"] show thesis
  13.716 +    by (intro that[of "-l"]) (simp add: ereal_uminus_eq_reorder)
  13.717  qed
  13.718  
  13.719 -lemma compact_extreal:
  13.720 -  fixes X :: "nat \<Rightarrow> extreal"
  13.721 +lemma compact_ereal:
  13.722 +  fixes X :: "nat \<Rightarrow> ereal"
  13.723    shows "\<exists>l r. subseq r \<and> (X \<circ> r) ----> l"
  13.724  proof -
  13.725    obtain r where "subseq r" and mono: "monoseq (X \<circ> r)"
  13.726 @@ -576,66 +578,66 @@
  13.727    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)"
  13.728      by (auto simp add: monoseq_def)
  13.729    then obtain l where "(X\<circ>r) ----> l"
  13.730 -     using lim_extreal_increasing[of "X \<circ> r"] lim_extreal_decreasing[of "X \<circ> r"] by auto
  13.731 +     using lim_ereal_increasing[of "X \<circ> r"] lim_ereal_decreasing[of "X \<circ> r"] by auto
  13.732    then show ?thesis using `subseq r` by auto
  13.733  qed
  13.734  
  13.735 -lemma extreal_Sup_lim:
  13.736 -  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
  13.737 +lemma ereal_Sup_lim:
  13.738 +  assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)"
  13.739    shows "a \<le> Sup s"
  13.740 -by (metis Lim_bounded_extreal assms complete_lattice_class.Sup_upper)
  13.741 +by (metis Lim_bounded_ereal assms complete_lattice_class.Sup_upper)
  13.742  
  13.743 -lemma extreal_Inf_lim:
  13.744 -  assumes "\<And>n. b n \<in> s" "b ----> (a::extreal)"
  13.745 +lemma ereal_Inf_lim:
  13.746 +  assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)"
  13.747    shows "Inf s \<le> a"
  13.748 -by (metis Lim_bounded2_extreal assms complete_lattice_class.Inf_lower)
  13.749 +by (metis Lim_bounded2_ereal assms complete_lattice_class.Inf_lower)
  13.750  
  13.751 -lemma SUP_Lim_extreal:
  13.752 -  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" "X ----> l" shows "(SUP n. X n) = l"
  13.753 -proof (rule extreal_SUPI)
  13.754 +lemma SUP_Lim_ereal:
  13.755 +  fixes X :: "nat \<Rightarrow> ereal" assumes "incseq X" "X ----> l" shows "(SUP n. X n) = l"
  13.756 +proof (rule ereal_SUPI)
  13.757    fix n from assms show "X n \<le> l"
  13.758 -    by (intro incseq_le_extreal) (simp add: incseq_def)
  13.759 +    by (intro incseq_le_ereal) (simp add: incseq_def)
  13.760  next
  13.761    fix y assume "\<And>n. n \<in> UNIV \<Longrightarrow> X n \<le> y"
  13.762 -  with extreal_Sup_lim[OF _ `X ----> l`, of "{..y}"]
  13.763 +  with ereal_Sup_lim[OF _ `X ----> l`, of "{..y}"]
  13.764    show "l \<le> y" by auto
  13.765  qed
  13.766  
  13.767 -lemma LIMSEQ_extreal_SUPR:
  13.768 -  fixes X :: "nat \<Rightarrow> extreal" assumes "incseq X" shows "X ----> (SUP n. X n)"
  13.769 -proof (rule lim_extreal_increasing)
  13.770 +lemma LIMSEQ_ereal_SUPR:
  13.771 +  fixes X :: "nat \<Rightarrow> ereal" assumes "incseq X" shows "X ----> (SUP n. X n)"
  13.772 +proof (rule lim_ereal_increasing)
  13.773    fix n m :: nat assume "m \<le> n" then show "X m \<le> X n"
  13.774      using `incseq X` by (simp add: incseq_def)
  13.775  next
  13.776    fix l assume "X ----> l"
  13.777 -  with SUP_Lim_extreal[of X, OF assms this] show ?thesis by simp
  13.778 +  with SUP_Lim_ereal[of X, OF assms this] show ?thesis by simp
  13.779  qed
  13.780  
  13.781 -lemma INF_Lim_extreal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::extreal)"
  13.782 -  using SUP_Lim_extreal[of "\<lambda>i. - X i" "- l"]
  13.783 -  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
  13.784 +lemma INF_Lim_ereal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::ereal)"
  13.785 +  using SUP_Lim_ereal[of "\<lambda>i. - X i" "- l"]
  13.786 +  by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
  13.787  
  13.788 -lemma LIMSEQ_extreal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: extreal)"
  13.789 -  using LIMSEQ_extreal_SUPR[of "\<lambda>i. - X i"]
  13.790 -  by (simp add: extreal_SUPR_uminus extreal_lim_uminus)
  13.791 +lemma LIMSEQ_ereal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: ereal)"
  13.792 +  using LIMSEQ_ereal_SUPR[of "\<lambda>i. - X i"]
  13.793 +  by (simp add: ereal_SUPR_uminus ereal_lim_uminus)
  13.794  
  13.795  lemma SUP_eq_LIMSEQ:
  13.796    assumes "mono f"
  13.797 -  shows "(SUP n. extreal (f n)) = extreal x \<longleftrightarrow> f ----> x"
  13.798 +  shows "(SUP n. ereal (f n)) = ereal x \<longleftrightarrow> f ----> x"
  13.799  proof
  13.800 -  have inc: "incseq (\<lambda>i. extreal (f i))"
  13.801 +  have inc: "incseq (\<lambda>i. ereal (f i))"
  13.802      using `mono f` unfolding mono_def incseq_def by auto
  13.803    { assume "f ----> x"
  13.804 -   then have "(\<lambda>i. extreal (f i)) ----> extreal x" by auto
  13.805 -   from SUP_Lim_extreal[OF inc this]
  13.806 -   show "(SUP n. extreal (f n)) = extreal x" . }
  13.807 -  { assume "(SUP n. extreal (f n)) = extreal x"
  13.808 -    with LIMSEQ_extreal_SUPR[OF inc]
  13.809 +   then have "(\<lambda>i. ereal (f i)) ----> ereal x" by auto
  13.810 +   from SUP_Lim_ereal[OF inc this]
  13.811 +   show "(SUP n. ereal (f n)) = ereal x" . }
  13.812 +  { assume "(SUP n. ereal (f n)) = ereal x"
  13.813 +    with LIMSEQ_ereal_SUPR[OF inc]
  13.814      show "f ----> x" by auto }
  13.815  qed
  13.816  
  13.817  lemma Liminf_within:
  13.818 -  fixes f :: "'a::metric_space => extreal"
  13.819 +  fixes f :: "'a::metric_space => ereal"
  13.820    shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
  13.821  proof-
  13.822  let ?l="(SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
  13.823 @@ -645,7 +647,7 @@
  13.824    { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
  13.825    moreover
  13.826    { assume "~(T=UNIV)"
  13.827 -    then obtain B where "T={B<..}" using T_def extreal_open_mono_set[of T] by auto
  13.828 +    then obtain B where "T={B<..}" using T_def ereal_open_mono_set[of T] by auto
  13.829      hence "B<?l" using T_def by auto
  13.830      then obtain d where d_def: "0<d & B<(INF y:(S Int ball x d - {x}). f y)"
  13.831        unfolding less_SUP_iff by auto
  13.832 @@ -670,14 +672,14 @@
  13.833      } hence "B <= INFI (S Int ball x d - {x}) f" apply (subst le_INFI) by auto
  13.834      also have "...<=?l" apply (subst le_SUPI) using d_def by auto
  13.835      finally have "B<=?l" by auto
  13.836 -  } hence "z <= ?l" using extreal_le_extreal[of z "?l"] by auto
  13.837 +  } hence "z <= ?l" using ereal_le_ereal[of z "?l"] by auto
  13.838  }
  13.839 -ultimately show ?thesis unfolding extreal_Liminf_Sup_monoset eventually_within
  13.840 -   apply (subst extreal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"]) by auto
  13.841 +ultimately show ?thesis unfolding ereal_Liminf_Sup_monoset eventually_within
  13.842 +   apply (subst ereal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"]) by auto
  13.843  qed
  13.844  
  13.845  lemma Limsup_within:
  13.846 -  fixes f :: "'a::metric_space => extreal"
  13.847 +  fixes f :: "'a::metric_space => ereal"
  13.848    shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
  13.849  proof-
  13.850  let ?l="(INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
  13.851 @@ -687,12 +689,12 @@
  13.852    { assume "T=UNIV" hence ?thesis by (simp add: gt_ex) }
  13.853    moreover
  13.854    { assume "~(T=UNIV)" hence "~(uminus ` T = UNIV)"
  13.855 -       by (metis Int_UNIV_right Int_absorb1 image_mono extreal_minus_minus_image subset_UNIV)
  13.856 -    hence "uminus ` T = {Inf (uminus ` T)<..}" using T_def extreal_open_mono_set[of "uminus ` T"]
  13.857 -       extreal_open_uminus[of T] by auto
  13.858 +       by (metis Int_UNIV_right Int_absorb1 image_mono ereal_minus_minus_image subset_UNIV)
  13.859 +    hence "uminus ` T = {Inf (uminus ` T)<..}" using T_def ereal_open_mono_set[of "uminus ` T"]
  13.860 +       ereal_open_uminus[of T] by auto
  13.861      then obtain B where "T={..<B}"
  13.862 -      unfolding extreal_Inf_uminus_image_eq extreal_uminus_lessThan[symmetric]
  13.863 -      unfolding inj_image_eq_iff[OF extreal_inj_on_uminus] by simp
  13.864 +      unfolding ereal_Inf_uminus_image_eq ereal_uminus_lessThan[symmetric]
  13.865 +      unfolding inj_image_eq_iff[OF ereal_inj_on_uminus] by simp
  13.866      hence "?l<B" using T_def by auto
  13.867      then obtain d where d_def: "0<d & (SUP y:(S Int ball x d - {x}). f y)<B"
  13.868        unfolding INF_less_iff by auto
  13.869 @@ -717,33 +719,33 @@
  13.870      } hence "SUPR (S Int ball x d - {x}) f <= B" apply (subst SUP_leI) by auto
  13.871      moreover have "?l<=SUPR (S Int ball x d - {x}) f" apply (subst INF_leI) using d_def by auto
  13.872      ultimately have "?l<=B" by auto
  13.873 -  } hence "?l <= z" using extreal_ge_extreal[of z "?l"] by auto
  13.874 +  } hence "?l <= z" using ereal_ge_ereal[of z "?l"] by auto
  13.875  }
  13.876 -ultimately show ?thesis unfolding extreal_Limsup_Inf_monoset eventually_within
  13.877 -   apply (subst extreal_InfI) by auto
  13.878 +ultimately show ?thesis unfolding ereal_Limsup_Inf_monoset eventually_within
  13.879 +   apply (subst ereal_InfI) by auto
  13.880  qed
  13.881  
  13.882  
  13.883  lemma Liminf_within_UNIV:
  13.884 -  fixes f :: "'a::metric_space => extreal"
  13.885 +  fixes f :: "'a::metric_space => ereal"
  13.886    shows "Liminf (at x) f = Liminf (at x within UNIV) f"
  13.887  by (metis within_UNIV)
  13.888  
  13.889  
  13.890  lemma Liminf_at:
  13.891 -  fixes f :: "'a::metric_space => extreal"
  13.892 +  fixes f :: "'a::metric_space => ereal"
  13.893    shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
  13.894  using Liminf_within[of x UNIV f] Liminf_within_UNIV[of x f] by auto
  13.895  
  13.896  
  13.897  lemma Limsup_within_UNIV:
  13.898 -  fixes f :: "'a::metric_space => extreal"
  13.899 +  fixes f :: "'a::metric_space => ereal"
  13.900    shows "Limsup (at x) f = Limsup (at x within UNIV) f"
  13.901  by (metis within_UNIV)
  13.902  
  13.903  
  13.904  lemma Limsup_at:
  13.905 -  fixes f :: "'a::metric_space => extreal"
  13.906 +  fixes f :: "'a::metric_space => ereal"
  13.907    shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
  13.908  using Limsup_within[of x UNIV f] Limsup_within_UNIV[of x f] by auto
  13.909  
  13.910 @@ -755,14 +757,14 @@
  13.911  by (metis assms(1) linorder_le_less_linear n_not_Suc_n real_of_nat_le_zero_cancel_iff)
  13.912  
  13.913  lemma Liminf_within_constant:
  13.914 -  fixes f :: "'a::metric_space => extreal"
  13.915 +  fixes f :: "'a::metric_space => ereal"
  13.916    assumes "ALL y:S. f y = C"
  13.917    assumes "~trivial_limit (at x within S)"
  13.918    shows "Liminf (at x within S) f = C"
  13.919  by (metis Lim_within_constant assms lim_imp_Liminf)
  13.920  
  13.921  lemma Limsup_within_constant:
  13.922 -  fixes f :: "'a::metric_space => extreal"
  13.923 +  fixes f :: "'a::metric_space => ereal"
  13.924    assumes "ALL y:S. f y = C"
  13.925    assumes "~trivial_limit (at x within S)"
  13.926    shows "Limsup (at x within S) f = C"
  13.927 @@ -805,17 +807,17 @@
  13.928  } ultimately show ?thesis by auto
  13.929  qed
  13.930  
  13.931 -lemma liminf_extreal_cminus:
  13.932 -  fixes f :: "nat \<Rightarrow> extreal" assumes "c \<noteq> -\<infinity>"
  13.933 +lemma liminf_ereal_cminus:
  13.934 +  fixes f :: "nat \<Rightarrow> ereal" assumes "c \<noteq> -\<infinity>"
  13.935    shows "liminf (\<lambda>x. c - f x) = c - limsup f"
  13.936  proof (cases c)
  13.937    case PInf then show ?thesis by (simp add: Liminf_const)
  13.938  next
  13.939    case (real r) then show ?thesis
  13.940      unfolding liminf_SUPR_INFI limsup_INFI_SUPR
  13.941 -    apply (subst INFI_extreal_cminus)
  13.942 +    apply (subst INFI_ereal_cminus)
  13.943      apply auto
  13.944 -    apply (subst SUPR_extreal_cminus)
  13.945 +    apply (subst SUPR_ereal_cminus)
  13.946      apply auto
  13.947      done
  13.948  qed (insert `c \<noteq> -\<infinity>`, simp)
  13.949 @@ -853,77 +855,77 @@
  13.950  from this show ?thesis using continuous_imp_tendsto by auto
  13.951  qed
  13.952  
  13.953 -lemma continuous_at_of_extreal:
  13.954 -  fixes x0 :: extreal
  13.955 +lemma continuous_at_of_ereal:
  13.956 +  fixes x0 :: ereal
  13.957    assumes "\<bar>x0\<bar> \<noteq> \<infinity>"
  13.958    shows "continuous (at x0) real"
  13.959  proof-
  13.960  { fix T assume T_def: "open T & real x0 : T"
  13.961 -  def S == "extreal ` T"
  13.962 -  hence "extreal (real x0) : S" using T_def by auto
  13.963 -  hence "x0 : S" using assms extreal_real by auto
  13.964 -  moreover have "open S" using open_extreal S_def T_def by auto
  13.965 +  def S == "ereal ` T"
  13.966 +  hence "ereal (real x0) : S" using T_def by auto
  13.967 +  hence "x0 : S" using assms ereal_real by auto
  13.968 +  moreover have "open S" using open_ereal S_def T_def by auto
  13.969    moreover have "ALL y:S. real y : T" using S_def T_def by auto
  13.970    ultimately have "EX S. x0 : S & open S & (ALL y:S. real y : T)" by auto
  13.971  } from this show ?thesis unfolding continuous_at_open by blast
  13.972  qed
  13.973  
  13.974  
  13.975 -lemma continuous_at_iff_extreal:
  13.976 +lemma continuous_at_iff_ereal:
  13.977  fixes f :: "'a::t2_space => real"
  13.978 -shows "continuous (at x0) f <-> continuous (at x0) (extreal o f)"
  13.979 +shows "continuous (at x0) f <-> continuous (at x0) (ereal o f)"
  13.980  proof-
  13.981 -{ assume "continuous (at x0) f" hence "continuous (at x0) (extreal o f)"
  13.982 -     using continuous_at_extreal continuous_at_compose[of x0 f extreal] by auto
  13.983 +{ assume "continuous (at x0) f" hence "continuous (at x0) (ereal o f)"
  13.984 +     using continuous_at_ereal continuous_at_compose[of x0 f ereal] by auto
  13.985  }
  13.986  moreover
  13.987 -{ assume "continuous (at x0) (extreal o f)"
  13.988 -  hence "continuous (at x0) (real o (extreal o f))"
  13.989 -     using continuous_at_of_extreal by (intro continuous_at_compose[of x0 "extreal o f"]) auto
  13.990 -  moreover have "real o (extreal o f) = f" using real_extreal_id by (simp add: o_assoc)
  13.991 +{ assume "continuous (at x0) (ereal o f)"
  13.992 +  hence "continuous (at x0) (real o (ereal o f))"
  13.993 +     using continuous_at_of_ereal by (intro continuous_at_compose[of x0 "ereal o f"]) auto
  13.994 +  moreover have "real o (ereal o f) = f" using real_ereal_id by (simp add: o_assoc)
  13.995    ultimately have "continuous (at x0) f" by auto
  13.996  } ultimately show ?thesis by auto
  13.997  qed
  13.998  
  13.999  
 13.1000 -lemma continuous_on_iff_extreal:
 13.1001 +lemma continuous_on_iff_ereal:
 13.1002  fixes f :: "'a::t2_space => real"
 13.1003  fixes A assumes "open A"
 13.1004 -shows "continuous_on A f <-> continuous_on A (extreal o f)"
 13.1005 -   using continuous_at_iff_extreal assms by (auto simp add: continuous_on_eq_continuous_at)
 13.1006 +shows "continuous_on A f <-> continuous_on A (ereal o f)"
 13.1007 +   using continuous_at_iff_ereal assms by (auto simp add: continuous_on_eq_continuous_at)
 13.1008  
 13.1009  
 13.1010 -lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>)}) real"
 13.1011 -   using continuous_at_of_extreal continuous_on_eq_continuous_at open_image_extreal by auto
 13.1012 +lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>::ereal)}) real"
 13.1013 +   using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal by auto
 13.1014  
 13.1015  
 13.1016  lemma continuous_on_iff_real:
 13.1017 -  fixes f :: "'a::t2_space => extreal"
 13.1018 +  fixes f :: "'a::t2_space => ereal"
 13.1019    assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
 13.1020    shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)"
 13.1021  proof-
 13.1022    have "f ` A <= UNIV-{\<infinity>,(-\<infinity>)}" using assms by force
 13.1023    hence *: "continuous_on (f ` A) real"
 13.1024       using continuous_on_real by (simp add: continuous_on_subset)
 13.1025 -have **: "continuous_on ((real o f) ` A) extreal"
 13.1026 -   using continuous_on_extreal continuous_on_subset[of "UNIV" "extreal" "(real o f) ` A"] by blast
 13.1027 +have **: "continuous_on ((real o f) ` A) ereal"
 13.1028 +   using continuous_on_ereal continuous_on_subset[of "UNIV" "ereal" "(real o f) ` A"] by blast
 13.1029  { assume "continuous_on A f" hence "continuous_on A (real o f)"
 13.1030    apply (subst continuous_on_compose) using * by auto
 13.1031  }
 13.1032  moreover
 13.1033  { assume "continuous_on A (real o f)"
 13.1034 -  hence "continuous_on A (extreal o (real o f))"
 13.1035 +  hence "continuous_on A (ereal o (real o f))"
 13.1036       apply (subst continuous_on_compose) using ** by auto
 13.1037    hence "continuous_on A f"
 13.1038 -     apply (subst continuous_on_eq[of A "extreal o (real o f)" f])
 13.1039 -     using assms extreal_real by auto
 13.1040 +     apply (subst continuous_on_eq[of A "ereal o (real o f)" f])
 13.1041 +     using assms ereal_real by auto
 13.1042  }
 13.1043  ultimately show ?thesis by auto
 13.1044  qed
 13.1045  
 13.1046  
 13.1047  lemma continuous_at_const:
 13.1048 -  fixes f :: "'a::t2_space => extreal"
 13.1049 +  fixes f :: "'a::t2_space => ereal"
 13.1050    assumes "ALL x. (f x = C)"
 13.1051    shows "ALL x. continuous (at x) f"
 13.1052  unfolding continuous_at_open using assms t1_space by auto
 13.1053 @@ -977,11 +979,11 @@
 13.1054  qed
 13.1055  
 13.1056  
 13.1057 -lemma mono_closed_extreal:
 13.1058 +lemma mono_closed_ereal:
 13.1059    fixes S :: "real set"
 13.1060    assumes mono: "ALL y z. y:S & y<=z --> z:S"
 13.1061    assumes "closed S"
 13.1062 -  shows "EX a. S = {x. a <= extreal x}"
 13.1063 +  shows "EX a. S = {x. a <= ereal x}"
 13.1064  proof-
 13.1065  { assume "S = {}" hence ?thesis apply(rule_tac x=PInfty in exI) by auto }
 13.1066  moreover
 13.1067 @@ -989,19 +991,21 @@
 13.1068  moreover
 13.1069  { assume "EX a. S = {a ..}"
 13.1070    from this obtain a where "S={a ..}" by auto
 13.1071 -  hence ?thesis apply(rule_tac x="extreal a" in exI) by auto
 13.1072 +  hence ?thesis apply(rule_tac x="ereal a" in exI) by auto
 13.1073  } ultimately show ?thesis using mono_closed_real[of S] assms by auto
 13.1074  qed
 13.1075  
 13.1076  subsection {* Sums *}
 13.1077  
 13.1078 -lemma setsum_extreal[simp]:
 13.1079 -  "(\<Sum>x\<in>A. extreal (f x)) = extreal (\<Sum>x\<in>A. f x)"
 13.1080 +lemma setsum_ereal[simp]:
 13.1081 +  "(\<Sum>x\<in>A. ereal (f x)) = ereal (\<Sum>x\<in>A. f x)"
 13.1082  proof cases
 13.1083    assume "finite A" then show ?thesis by induct auto
 13.1084  qed simp
 13.1085  
 13.1086 -lemma setsum_Pinfty: "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))"
 13.1087 +lemma setsum_Pinfty:
 13.1088 +  fixes f :: "'a \<Rightarrow> ereal"
 13.1089 +  shows "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))"
 13.1090  proof safe
 13.1091    assume *: "setsum f P = \<infinity>"
 13.1092    show "finite P"
 13.1093 @@ -1023,15 +1027,16 @@
 13.1094  qed
 13.1095  
 13.1096  lemma setsum_Inf:
 13.1097 +  fixes f :: "'a \<Rightarrow> ereal"
 13.1098    shows "\<bar>setsum f A\<bar> = \<infinity> \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>))"
 13.1099  proof
 13.1100    assume *: "\<bar>setsum f A\<bar> = \<infinity>"
 13.1101    have "finite A" by (rule ccontr) (insert *, auto)
 13.1102    moreover have "\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>"
 13.1103    proof (rule ccontr)
 13.1104 -    assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
 13.1105 +    assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto
 13.1106      from bchoice[OF this] guess r ..
 13.1107 -    with * show False by (auto simp: setsum_extreal)
 13.1108 +    with * show False by (auto simp: setsum_ereal)
 13.1109    qed
 13.1110    ultimately show "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)" by auto
 13.1111  next
 13.1112 @@ -1040,72 +1045,73 @@
 13.1113    then show "\<bar>setsum f A\<bar> = \<infinity>"
 13.1114    proof induct
 13.1115      case (insert j A) then show ?case
 13.1116 -      by (cases rule: extreal3_cases[of "f i" "f j" "setsum f A"]) auto
 13.1117 +      by (cases rule: ereal3_cases[of "f i" "f j" "setsum f A"]) auto
 13.1118    qed simp
 13.1119  qed
 13.1120  
 13.1121 -lemma setsum_real_of_extreal:
 13.1122 +lemma setsum_real_of_ereal:
 13.1123 +  fixes f :: "'i \<Rightarrow> ereal"
 13.1124    assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
 13.1125    shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
 13.1126  proof -
 13.1127 -  have "\<forall>x\<in>S. \<exists>r. f x = extreal r"
 13.1128 +  have "\<forall>x\<in>S. \<exists>r. f x = ereal r"
 13.1129    proof
 13.1130      fix x assume "x \<in> S"
 13.1131 -    from assms[OF this] show "\<exists>r. f x = extreal r" by (cases "f x") auto
 13.1132 +    from assms[OF this] show "\<exists>r. f x = ereal r" by (cases "f x") auto
 13.1133    qed
 13.1134    from bchoice[OF this] guess r ..
 13.1135    then show ?thesis by simp
 13.1136  qed
 13.1137  
 13.1138 -lemma setsum_extreal_0:
 13.1139 -  fixes f :: "'a \<Rightarrow> extreal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
 13.1140 +lemma setsum_ereal_0:
 13.1141 +  fixes f :: "'a \<Rightarrow> ereal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
 13.1142    shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
 13.1143  proof
 13.1144    assume *: "(\<Sum>x\<in>A. f x) = 0"
 13.1145    then have "(\<Sum>x\<in>A. f x) \<noteq> \<infinity>" by auto
 13.1146    then have "\<forall>i\<in>A. \<bar>f i\<bar> \<noteq> \<infinity>" using assms by (force simp: setsum_Pinfty)
 13.1147 -  then have "\<forall>i\<in>A. \<exists>r. f i = extreal r" by auto
 13.1148 +  then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto
 13.1149    from bchoice[OF this] * assms show "\<forall>i\<in>A. f i = 0"
 13.1150      using setsum_nonneg_eq_0_iff[of A "\<lambda>i. real (f i)"] by auto
 13.1151  qed (rule setsum_0')
 13.1152  
 13.1153  
 13.1154 -lemma setsum_extreal_right_distrib:
 13.1155 -  fixes f :: "'a \<Rightarrow> extreal" assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
 13.1156 +lemma setsum_ereal_right_distrib:
 13.1157 +  fixes f :: "'a \<Rightarrow> ereal" assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
 13.1158    shows "r * setsum f A = (\<Sum>n\<in>A. r * f n)"
 13.1159  proof cases
 13.1160    assume "finite A" then show ?thesis using assms
 13.1161 -    by induct (auto simp: extreal_right_distrib setsum_nonneg)
 13.1162 +    by induct (auto simp: ereal_right_distrib setsum_nonneg)
 13.1163  qed simp
 13.1164  
 13.1165 -lemma sums_extreal_positive:
 13.1166 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
 13.1167 +lemma sums_ereal_positive:
 13.1168 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" shows "f sums (SUP n. \<Sum>i<n. f i)"
 13.1169  proof -
 13.1170    have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)"
 13.1171 -    using extreal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
 13.1172 -  from LIMSEQ_extreal_SUPR[OF this]
 13.1173 +    using ereal_add_mono[OF _ assms] by (auto intro!: incseq_SucI)
 13.1174 +  from LIMSEQ_ereal_SUPR[OF this]
 13.1175    show ?thesis unfolding sums_def by (simp add: atLeast0LessThan)
 13.1176  qed
 13.1177  
 13.1178 -lemma summable_extreal_pos:
 13.1179 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i" shows "summable f"
 13.1180 -  using sums_extreal_positive[of f, OF assms] unfolding summable_def by auto
 13.1181 +lemma summable_ereal_pos:
 13.1182 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i" shows "summable f"
 13.1183 +  using sums_ereal_positive[of f, OF assms] unfolding summable_def by auto
 13.1184  
 13.1185 -lemma suminf_extreal_eq_SUPR:
 13.1186 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>i. 0 \<le> f i"
 13.1187 +lemma suminf_ereal_eq_SUPR:
 13.1188 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>i. 0 \<le> f i"
 13.1189    shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)"
 13.1190 -  using sums_extreal_positive[of f, OF assms, THEN sums_unique] by simp
 13.1191 +  using sums_ereal_positive[of f, OF assms, THEN sums_unique] by simp
 13.1192  
 13.1193 -lemma sums_extreal:
 13.1194 -  "(\<lambda>x. extreal (f x)) sums extreal x \<longleftrightarrow> f sums x"
 13.1195 +lemma sums_ereal:
 13.1196 +  "(\<lambda>x. ereal (f x)) sums ereal x \<longleftrightarrow> f sums x"
 13.1197    unfolding sums_def by simp
 13.1198  
 13.1199  lemma suminf_bound:
 13.1200 -  fixes f :: "nat \<Rightarrow> extreal"
 13.1201 +  fixes f :: "nat \<Rightarrow> ereal"
 13.1202    assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x" and pos: "\<And>n. 0 \<le> f n"
 13.1203    shows "suminf f \<le> x"
 13.1204 -proof (rule Lim_bounded_extreal)
 13.1205 -  have "summable f" using pos[THEN summable_extreal_pos] .
 13.1206 +proof (rule Lim_bounded_ereal)
 13.1207 +  have "summable f" using pos[THEN summable_ereal_pos] .
 13.1208    then show "(\<lambda>N. \<Sum>n<N. f n) ----> suminf f"
 13.1209      by (auto dest!: summable_sums simp: sums_def atLeast0LessThan)
 13.1210    show "\<forall>n\<ge>0. setsum f {..<n} \<le> x"
 13.1211 @@ -1113,15 +1119,15 @@
 13.1212  qed
 13.1213  
 13.1214  lemma suminf_bound_add:
 13.1215 -  fixes f :: "nat \<Rightarrow> extreal"
 13.1216 +  fixes f :: "nat \<Rightarrow> ereal"
 13.1217    assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x" and pos: "\<And>n. 0 \<le> f n" and "y \<noteq> -\<infinity>"
 13.1218    shows "suminf f + y \<le> x"
 13.1219  proof (cases y)
 13.1220    case (real r) then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y"
 13.1221 -    using assms by (simp add: extreal_le_minus)
 13.1222 +    using assms by (simp add: ereal_le_minus)
 13.1223    then have "(\<Sum> n. f n) \<le> x - y" using pos by (rule suminf_bound)
 13.1224    then show "(\<Sum> n. f n) + y \<le> x"
 13.1225 -    using assms real by (simp add: extreal_le_minus)
 13.1226 +    using assms real by (simp add: ereal_le_minus)
 13.1227  qed (insert assms, auto)
 13.1228  
 13.1229  lemma sums_finite:
 13.1230 @@ -1140,22 +1146,22 @@
 13.1231    shows "suminf f = (\<Sum>N<n. f N)"
 13.1232    using sums_finite[OF assms, THEN sums_unique] by simp
 13.1233  
 13.1234 -lemma suminf_extreal_0[simp]: "(\<Sum>i. 0) = (0::'a::{comm_monoid_add,t2_space})"
 13.1235 +lemma suminf_ereal_0[simp]: "(\<Sum>i. 0) = (0::'a::{comm_monoid_add,t2_space})"
 13.1236    using suminf_finite[of 0 "\<lambda>x. 0"] by simp
 13.1237  
 13.1238  lemma suminf_upper:
 13.1239 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
 13.1240 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>n. 0 \<le> f n"
 13.1241    shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)"
 13.1242 -  unfolding suminf_extreal_eq_SUPR[OF assms] SUPR_def
 13.1243 +  unfolding suminf_ereal_eq_SUPR[OF assms] SUPR_def
 13.1244    by (auto intro: complete_lattice_class.Sup_upper image_eqI)
 13.1245  
 13.1246  lemma suminf_0_le:
 13.1247 -  fixes f :: "nat \<Rightarrow> extreal" assumes "\<And>n. 0 \<le> f n"
 13.1248 +  fixes f :: "nat \<Rightarrow> ereal" assumes "\<And>n. 0 \<le> f n"
 13.1249    shows "0 \<le> (\<Sum>n. f n)"
 13.1250    using suminf_upper[of f 0, OF assms] by simp
 13.1251  
 13.1252  lemma suminf_le_pos:
 13.1253 -  fixes f g :: "nat \<Rightarrow> extreal"
 13.1254 +  fixes f g :: "nat \<Rightarrow> ereal"
 13.1255    assumes "\<And>N. f N \<le> g N" "\<And>N. 0 \<le> f N"
 13.1256    shows "suminf f \<le> suminf g"
 13.1257  proof (safe intro!: suminf_bound)
 13.1258 @@ -1165,27 +1171,28 @@
 13.1259    finally show "setsum f {..<n} \<le> suminf g" .
 13.1260  qed (rule assms(2))
 13.1261  
 13.1262 -lemma suminf_half_series_extreal: "(\<Sum>n. (1/2 :: extreal)^Suc n) = 1"
 13.1263 -  using sums_extreal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]
 13.1264 -  by (simp add: one_extreal_def)
 13.1265 +lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal)^Suc n) = 1"
 13.1266 +  using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]
 13.1267 +  by (simp add: one_ereal_def)
 13.1268  
 13.1269 -lemma suminf_add_extreal:
 13.1270 -  fixes f g :: "nat \<Rightarrow> extreal"
 13.1271 +lemma suminf_add_ereal:
 13.1272 +  fixes f g :: "nat \<Rightarrow> ereal"
 13.1273    assumes "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i"
 13.1274    shows "(\<Sum>i. f i + g i) = suminf f + suminf g"
 13.1275 -  apply (subst (1 2 3) suminf_extreal_eq_SUPR)
 13.1276 +  apply (subst (1 2 3) suminf_ereal_eq_SUPR)
 13.1277    unfolding setsum_addf
 13.1278 -  by (intro assms extreal_add_nonneg_nonneg SUPR_extreal_add_pos incseq_setsumI setsum_nonneg ballI)+
 13.1279 +  by (intro assms ereal_add_nonneg_nonneg SUPR_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+
 13.1280  
 13.1281 -lemma suminf_cmult_extreal:
 13.1282 -  fixes f g :: "nat \<Rightarrow> extreal"
 13.1283 +lemma suminf_cmult_ereal:
 13.1284 +  fixes f g :: "nat \<Rightarrow> ereal"
 13.1285    assumes "\<And>i. 0 \<le> f i" "0 \<le> a"
 13.1286    shows "(\<Sum>i. a * f i) = a * suminf f"
 13.1287 -  by (auto simp: setsum_extreal_right_distrib[symmetric] assms
 13.1288 -                 extreal_zero_le_0_iff setsum_nonneg suminf_extreal_eq_SUPR
 13.1289 -           intro!: SUPR_extreal_cmult )
 13.1290 +  by (auto simp: setsum_ereal_right_distrib[symmetric] assms
 13.1291 +                 ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUPR
 13.1292 +           intro!: SUPR_ereal_cmult )
 13.1293  
 13.1294  lemma suminf_PInfty:
 13.1295 +  fixes f :: "nat \<Rightarrow> ereal"
 13.1296    assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
 13.1297    shows "f i \<noteq> \<infinity>"
 13.1298  proof -
 13.1299 @@ -1197,43 +1204,43 @@
 13.1300  
 13.1301  lemma suminf_PInfty_fun:
 13.1302    assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>"
 13.1303 -  shows "\<exists>f'. f = (\<lambda>x. extreal (f' x))"
 13.1304 +  shows "\<exists>f'. f = (\<lambda>x. ereal (f' x))"
 13.1305  proof -
 13.1306 -  have "\<forall>i. \<exists>r. f i = extreal r"
 13.1307 +  have "\<forall>i. \<exists>r. f i = ereal r"
 13.1308    proof
 13.1309 -    fix i show "\<exists>r. f i = extreal r"
 13.1310 +    fix i show "\<exists>r. f i = ereal r"
 13.1311        using suminf_PInfty[OF assms] assms(1)[of i] by (cases "f i") auto
 13.1312    qed
 13.1313    from choice[OF this] show ?thesis by auto
 13.1314  qed
 13.1315  
 13.1316 -lemma summable_extreal:
 13.1317 -  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
 13.1318 +lemma summable_ereal:
 13.1319 +  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"
 13.1320    shows "summable f"
 13.1321  proof -
 13.1322 -  have "0 \<le> (\<Sum>i. extreal (f i))"
 13.1323 +  have "0 \<le> (\<Sum>i. ereal (f i))"
 13.1324      using assms by (intro suminf_0_le) auto
 13.1325 -  with assms obtain r where r: "(\<Sum>i. extreal (f i)) = extreal r"
 13.1326 -    by (cases "\<Sum>i. extreal (f i)") auto
 13.1327 -  from summable_extreal_pos[of "\<lambda>x. extreal (f x)"]
 13.1328 -  have "summable (\<lambda>x. extreal (f x))" using assms by auto
 13.1329 +  with assms obtain r where r: "(\<Sum>i. ereal (f i)) = ereal r"
 13.1330 +    by (cases "\<Sum>i. ereal (f i)") auto
 13.1331 +  from summable_ereal_pos[of "\<lambda>x. ereal (f x)"]
 13.1332 +  have "summable (\<lambda>x. ereal (f x))" using assms by auto
 13.1333    from summable_sums[OF this]
 13.1334 -  have "(\<lambda>x. extreal (f x)) sums (\<Sum>x. extreal (f x))" by auto
 13.1335 +  have "(\<lambda>x. ereal (f x)) sums (\<Sum>x. ereal (f x))" by auto
 13.1336    then show "summable f"
 13.1337 -    unfolding r sums_extreal summable_def ..
 13.1338 +    unfolding r sums_ereal summable_def ..
 13.1339  qed
 13.1340  
 13.1341 -lemma suminf_extreal:
 13.1342 -  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. extreal (f i)) \<noteq> \<infinity>"
 13.1343 -  shows "(\<Sum>i. extreal (f i)) = extreal (suminf f)"
 13.1344 +lemma suminf_ereal:
 13.1345 +  assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"
 13.1346 +  shows "(\<Sum>i. ereal (f i)) = ereal (suminf f)"
 13.1347  proof (rule sums_unique[symmetric])
 13.1348 -  from summable_extreal[OF assms]
 13.1349 -  show "(\<lambda>x. extreal (f x)) sums (extreal (suminf f))"
 13.1350 -    unfolding sums_extreal using assms by (intro summable_sums summable_extreal)
 13.1351 +  from summable_ereal[OF assms]
 13.1352 +  show "(\<lambda>x. ereal (f x)) sums (ereal (suminf f))"
 13.1353 +    unfolding sums_ereal using assms by (intro summable_sums summable_ereal)
 13.1354  qed
 13.1355  
 13.1356 -lemma suminf_extreal_minus:
 13.1357 -  fixes f g :: "nat \<Rightarrow> extreal"
 13.1358 +lemma suminf_ereal_minus:
 13.1359 +  fixes f g :: "nat \<Rightarrow> ereal"
 13.1360    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>"
 13.1361    shows "(\<Sum>i. f i - g i) = suminf f - suminf g"
 13.1362  proof -
 13.1363 @@ -1241,50 +1248,51 @@
 13.1364    moreover
 13.1365    from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] guess f' .. note this[simp]
 13.1366    from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] guess g' .. note this[simp]
 13.1367 -  { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: extreal_le_minus_iff) }
 13.1368 +  { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: ereal_le_minus_iff) }
 13.1369    moreover
 13.1370    have "suminf (\<lambda>i. f i - g i) \<le> suminf f"
 13.1371      using assms by (auto intro!: suminf_le_pos simp: field_simps)
 13.1372    then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>" using fin by auto
 13.1373    ultimately show ?thesis using assms `\<And>i. 0 \<le> f i`
 13.1374      apply simp
 13.1375 -    by (subst (1 2 3) suminf_extreal)
 13.1376 -       (auto intro!: suminf_diff[symmetric] summable_extreal)
 13.1377 +    by (subst (1 2 3) suminf_ereal)
 13.1378 +       (auto intro!: suminf_diff[symmetric] summable_ereal)
 13.1379  qed
 13.1380  
 13.1381 -lemma suminf_extreal_PInf[simp]:
 13.1382 -  "(\<Sum>x. \<infinity>) = \<infinity>"
 13.1383 +lemma suminf_ereal_PInf[simp]:
 13.1384 +  "(\<Sum>x. \<infinity>::ereal) = \<infinity>"
 13.1385  proof -
 13.1386 -  have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>)" by (rule suminf_upper) auto
 13.1387 +  have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>::ereal)" by (rule suminf_upper) auto
 13.1388    then show ?thesis by simp
 13.1389  qed
 13.1390  
 13.1391 -lemma summable_real_of_extreal:
 13.1392 +lemma summable_real_of_ereal:
 13.1393 +  fixes f :: "nat \<Rightarrow> ereal"
 13.1394    assumes f: "\<And>i. 0 \<le> f i" and fin: "(\<Sum>i. f i) \<noteq> \<infinity>"
 13.1395    shows "summable (\<lambda>i. real (f i))"
 13.1396  proof (rule summable_def[THEN iffD2])
 13.1397    have "0 \<le> (\<Sum>i. f i)" using assms by (auto intro: suminf_0_le)
 13.1398 -  with fin obtain r where r: "extreal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto
 13.1399 +  with fin obtain r where r: "ereal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto
 13.1400    { fix i have "f i \<noteq> \<infinity>" using f by (intro suminf_PInfty[OF _ fin]) auto
 13.1401      then have "\<bar>f i\<bar> \<noteq> \<infinity>" using f[of i] by auto }
 13.1402    note fin = this
 13.1403 -  have "(\<lambda>i. extreal (real (f i))) sums (\<Sum>i. extreal (real (f i)))"
 13.1404 -    using f by (auto intro!: summable_extreal_pos summable_sums simp: extreal_le_real_iff zero_extreal_def)
 13.1405 -  also have "\<dots> = extreal r" using fin r by (auto simp: extreal_real)
 13.1406 -  finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_extreal)
 13.1407 +  have "(\<lambda>i. ereal (real (f i))) sums (\<Sum>i. ereal (real (f i)))"
 13.1408 +    using f by (auto intro!: summable_ereal_pos summable_sums simp: ereal_le_real_iff zero_ereal_def)
 13.1409 +  also have "\<dots> = ereal r" using fin r by (auto simp: ereal_real)
 13.1410 +  finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_ereal)
 13.1411  qed
 13.1412  
 13.1413  lemma suminf_SUP_eq:
 13.1414 -  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> extreal"
 13.1415 +  fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ereal"
 13.1416    assumes "\<And>i. incseq (\<lambda>n. f n i)" "\<And>n i. 0 \<le> f n i"
 13.1417    shows "(\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)"
 13.1418  proof -
 13.1419    { fix n :: nat
 13.1420      have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
 13.1421 -      using assms by (auto intro!: SUPR_extreal_setsum[symmetric]) }
 13.1422 +      using assms by (auto intro!: SUPR_ereal_setsum[symmetric]) }
 13.1423    note * = this
 13.1424    show ?thesis using assms
 13.1425 -    apply (subst (1 2) suminf_extreal_eq_SUPR)
 13.1426 +    apply (subst (1 2) suminf_ereal_eq_SUPR)
 13.1427      unfolding *
 13.1428      apply (auto intro!: le_SUPI2)
 13.1429      apply (subst SUP_commute) ..
    14.1 --- a/src/HOL/Probability/Binary_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    14.2 +++ b/src/HOL/Probability/Binary_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    14.3 @@ -291,7 +291,7 @@
    14.4          (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
    14.5        by (auto intro!: M2.measure_compl simp: vimage_Diff)
    14.6      with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
    14.7 -      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff)
    14.8 +      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_ereal_diff)
    14.9    next
   14.10      fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
   14.11      moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
   14.12 @@ -401,7 +401,7 @@
   14.13    apply (simp add: pair_measure_def pair_measure_generator_def)
   14.14  proof (rule M1.positive_integral_cong)
   14.15    fix x assume "x \<in> space M1"
   14.16 -  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)"
   14.17 +  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: ereal)"
   14.18      unfolding indicator_def by auto
   14.19    show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
   14.20      unfolding *
   14.21 @@ -656,7 +656,7 @@
   14.22    show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
   14.23      and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
   14.24      by (auto simp del: vimage_Int cong: measurable_cong
   14.25 -             intro!: M1.borel_measurable_extreal_setsum setsum_cong
   14.26 +             intro!: M1.borel_measurable_ereal_setsum setsum_cong
   14.27               simp add: M1.positive_integral_setsum simple_integral_def
   14.28                         M1.positive_integral_cmult
   14.29                         M1.positive_integral_cong[OF eq]
   14.30 @@ -760,7 +760,7 @@
   14.31      show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
   14.32        by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
   14.33      show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
   14.34 -      by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N)
   14.35 +      by (intro M1.borel_measurable_ereal_neq_const measure_cut_measurable_fst N)
   14.36      { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
   14.37        have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
   14.38        proof (rule M2.AE_I)
   14.39 @@ -822,45 +822,45 @@
   14.40    shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
   14.41      and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
   14.42  proof -
   14.43 -  let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)"
   14.44 +  let "?pf x" = "ereal (f x)" and "?nf x" = "ereal (- f x)"
   14.45    have
   14.46      borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
   14.47      int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
   14.48      using assms by auto
   14.49 -  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   14.50 -     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   14.51 +  have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   14.52 +     "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
   14.53      using borel[THEN positive_integral_fst_measurable(1)] int
   14.54      unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all
   14.55    with borel[THEN positive_integral_fst_measurable(1)]
   14.56 -  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   14.57 -    "AE x in M1. (\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   14.58 +  have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   14.59 +    "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
   14.60      by (auto intro!: M1.positive_integral_PInf_AE )
   14.61 -  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   14.62 -    "AE x in M1. \<bar>\<integral>\<^isup>+y. extreal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   14.63 +  then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   14.64 +    "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
   14.65      by (auto simp: M2.positive_integral_positive)
   14.66    from AE_pos show ?AE using assms
   14.67      by (simp add: measurable_pair_image_snd integrable_def)
   14.68 -  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   14.69 +  { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   14.70        using M2.positive_integral_positive
   14.71 -      by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder)
   14.72 -    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. extreal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
   14.73 +      by (intro M1.positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
   14.74 +    then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
   14.75    note this[simp]
   14.76 -  { fix f assume borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable P"
   14.77 -      and int: "integral\<^isup>P P (\<lambda>x. extreal (f x)) \<noteq> \<infinity>"
   14.78 -      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
   14.79 -    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. extreal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
   14.80 +  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable P"
   14.81 +      and int: "integral\<^isup>P P (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
   14.82 +      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
   14.83 +    have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
   14.84      proof (intro integrable_def[THEN iffD2] conjI)
   14.85        show "?f \<in> borel_measurable M1"
   14.86 -        using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable)
   14.87 -      have "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. extreal (f (x, y))  \<partial>M2) \<partial>M1)"
   14.88 +        using borel by (auto intro!: M1.borel_measurable_real_of_ereal positive_integral_fst_measurable)
   14.89 +      have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y))  \<partial>M2) \<partial>M1)"
   14.90          using AE M2.positive_integral_positive
   14.91 -        by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real)
   14.92 -      then show "(\<integral>\<^isup>+x. extreal (?f x) \<partial>M1) \<noteq> \<infinity>"
   14.93 +        by (auto intro!: M1.positive_integral_cong_AE simp: ereal_real)
   14.94 +      then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
   14.95          using positive_integral_fst_measurable[OF borel] int by simp
   14.96 -      have "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   14.97 +      have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
   14.98          by (intro M1.positive_integral_cong_pos)
   14.99 -           (simp add: M2.positive_integral_positive real_of_extreal_pos)
  14.100 -      then show "(\<integral>\<^isup>+x. extreal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
  14.101 +           (simp add: M2.positive_integral_positive real_of_ereal_pos)
  14.102 +      then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
  14.103      qed }
  14.104    with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
  14.105    show ?INT
    15.1 --- a/src/HOL/Probability/Borel_Space.thy	Wed Jul 20 10:11:08 2011 +0200
    15.2 +++ b/src/HOL/Probability/Borel_Space.thy	Wed Jul 20 10:48:00 2011 +0200
    15.3 @@ -112,7 +112,7 @@
    15.4  qed
    15.5  
    15.6  lemma (in sigma_algebra) borel_measurable_restricted:
    15.7 -  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
    15.8 +  fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M"
    15.9    shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
   15.10      (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
   15.11      (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
   15.12 @@ -123,7 +123,7 @@
   15.13    show ?thesis unfolding *
   15.14      unfolding in_borel_measurable_borel
   15.15    proof (simp, safe)
   15.16 -    fix S :: "extreal set" assume "S \<in> sets borel"
   15.17 +    fix S :: "ereal set" assume "S \<in> sets borel"
   15.18        "\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
   15.19      then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
   15.20      then have f: "?f -` S \<inter> A \<in> sets M"
   15.21 @@ -142,7 +142,7 @@
   15.22        then show ?thesis using f by auto
   15.23      qed
   15.24    next
   15.25 -    fix S :: "extreal set" assume "S \<in> sets borel"
   15.26 +    fix S :: "ereal set" assume "S \<in> sets borel"
   15.27        "\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M"
   15.28      then have f: "?f -` S \<inter> space M \<in> sets M" by auto
   15.29      then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
   15.30 @@ -1095,70 +1095,71 @@
   15.31  
   15.32  subsection "Borel space on the extended reals"
   15.33  
   15.34 -lemma borel_measurable_extreal_borel:
   15.35 -  "extreal \<in> borel_measurable borel"
   15.36 -  unfolding borel_def[where 'a=extreal]
   15.37 +lemma borel_measurable_ereal_borel:
   15.38 +  "ereal \<in> borel_measurable borel"
   15.39 +  unfolding borel_def[where 'a=ereal]
   15.40  proof (rule borel.measurable_sigma)
   15.41 -  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
   15.42 +  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
   15.43    then have "open X" by (auto simp: mem_def)
   15.44 -  then have "open (extreal -` X \<inter> space borel)"
   15.45 -    by (simp add: open_extreal_vimage)
   15.46 -  then show "extreal -` X \<inter> space borel \<in> sets borel" by auto
   15.47 +  then have "open (ereal -` X \<inter> space borel)"
   15.48 +    by (simp add: open_ereal_vimage)
   15.49 +  then show "ereal -` X \<inter> space borel \<in> sets borel" by auto
   15.50  qed auto
   15.51  
   15.52 -lemma (in sigma_algebra) borel_measurable_extreal[simp, intro]:
   15.53 -  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
   15.54 -  using measurable_comp[OF f borel_measurable_extreal_borel] unfolding comp_def .
   15.55 +lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]:
   15.56 +  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
   15.57 +  using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def .
   15.58  
   15.59 -lemma borel_measurable_real_of_extreal_borel:
   15.60 -  "(real :: extreal \<Rightarrow> real) \<in> borel_measurable borel"
   15.61 +lemma borel_measurable_real_of_ereal_borel:
   15.62 +  "(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel"
   15.63    unfolding borel_def[where 'a=real]
   15.64  proof (rule borel.measurable_sigma)
   15.65    fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = open \<rparr>"
   15.66    then have "open B" by (auto simp: mem_def)
   15.67 -  have *: "extreal -` real -` (B - {0}) = B - {0}" by auto
   15.68 -  have open_real: "open (real -` (B - {0}) :: extreal set)"
   15.69 -    unfolding open_extreal_def * using `open B` by auto
   15.70 -  show "(real -` B \<inter> space borel :: extreal set) \<in> sets borel"
   15.71 +  have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
   15.72 +  have open_real: "open (real -` (B - {0}) :: ereal set)"
   15.73 +    unfolding open_ereal_def * using `open B` by auto
   15.74 +  show "(real -` B \<inter> space borel :: ereal set) \<in> sets borel"
   15.75    proof cases
   15.76      assume "0 \<in> B"
   15.77 -    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0}"
   15.78 -      by (auto simp add: real_of_extreal_eq_0)
   15.79 -    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
   15.80 +    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0::ereal}"
   15.81 +      by (auto simp add: real_of_ereal_eq_0)
   15.82 +    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
   15.83        using open_real by auto
   15.84    next
   15.85      assume "0 \<notin> B"
   15.86 -    then have *: "(real -` B :: extreal set) = real -` (B - {0})"
   15.87 -      by (auto simp add: real_of_extreal_eq_0)
   15.88 -    then show "(real -` B :: extreal set) \<inter> space borel \<in> sets borel"
   15.89 +    then have *: "(real -` B :: ereal set) = real -` (B - {0})"
   15.90 +      by (auto simp add: real_of_ereal_eq_0)
   15.91 +    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
   15.92        using open_real by auto
   15.93    qed
   15.94  qed auto
   15.95  
   15.96 -lemma (in sigma_algebra) borel_measurable_real_of_extreal[simp, intro]:
   15.97 -  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: extreal)) \<in> borel_measurable M"
   15.98 -  using measurable_comp[OF f borel_measurable_real_of_extreal_borel] unfolding comp_def .
   15.99 +lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]:
  15.100 +  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M"
  15.101 +  using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def .
  15.102  
  15.103 -lemma (in sigma_algebra) borel_measurable_extreal_iff:
  15.104 -  shows "(\<lambda>x. extreal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
  15.105 +lemma (in sigma_algebra) borel_measurable_ereal_iff:
  15.106 +  shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
  15.107  proof
  15.108 -  assume "(\<lambda>x. extreal (f x)) \<in> borel_measurable M"
  15.109 -  from borel_measurable_real_of_extreal[OF this]
  15.110 +  assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
  15.111 +  from borel_measurable_real_of_ereal[OF this]
  15.112    show "f \<in> borel_measurable M" by auto
  15.113  qed auto
  15.114  
  15.115 -lemma (in sigma_algebra) borel_measurable_extreal_iff_real:
  15.116 -  "f \<in> borel_measurable M \<longleftrightarrow>
  15.117 +lemma (in sigma_algebra) borel_measurable_ereal_iff_real:
  15.118 +  fixes f :: "'a \<Rightarrow> ereal"
  15.119 +  shows "f \<in> borel_measurable M \<longleftrightarrow>
  15.120      ((\<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)"
  15.121  proof safe
  15.122    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"
  15.123    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
  15.124    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
  15.125 -  let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else extreal (real (f x))"
  15.126 +  let "?f x" = "if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real (f x))"
  15.127    have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
  15.128 -  also have "?f = f" by (auto simp: fun_eq_iff extreal_real)
  15.129 +  also have "?f = f" by (auto simp: fun_eq_iff ereal_real)
  15.130    finally show "f \<in> borel_measurable M" .
  15.131 -qed (auto intro: measurable_sets borel_measurable_real_of_extreal)
  15.132 +qed (auto intro: measurable_sets borel_measurable_real_of_ereal)
  15.133  
  15.134  lemma (in sigma_algebra) less_eq_ge_measurable:
  15.135    fixes f :: "'a \<Rightarrow> 'c::linorder"
  15.136 @@ -1186,95 +1187,96 @@
  15.137    ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
  15.138  qed
  15.139  
  15.140 -lemma (in sigma_algebra) borel_measurable_uminus_borel_extreal:
  15.141 -  "(uminus :: extreal \<Rightarrow> extreal) \<in> borel_measurable borel"
  15.142 +lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal:
  15.143 +  "(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel"
  15.144  proof (subst borel_def, rule borel.measurable_sigma)
  15.145 -  fix X :: "extreal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
  15.146 +  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = open\<rparr>"
  15.147    then have "open X" by (simp add: mem_def)
  15.148    have "uminus -` X = uminus ` X" by (force simp: image_iff)
  15.149 -  then have "open (uminus -` X)" using `open X` extreal_open_uminus by auto
  15.150 +  then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto
  15.151    then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
  15.152  qed auto
  15.153  
  15.154 -lemma (in sigma_algebra) borel_measurable_uminus_extreal[intro]:
  15.155 +lemma (in sigma_algebra) borel_measurable_uminus_ereal[intro]:
  15.156    assumes "f \<in> borel_measurable M"
  15.157 -  shows "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M"
  15.158 -  using measurable_comp[OF assms borel_measurable_uminus_borel_extreal] by (simp add: comp_def)
  15.159 +  shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
  15.160 +  using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def)
  15.161  
  15.162 -lemma (in sigma_algebra) borel_measurable_uminus_eq_extreal[simp]:
  15.163 -  "(\<lambda>x. - f x :: extreal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
  15.164 +lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]:
  15.165 +  "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
  15.166  proof
  15.167 -  assume ?l from borel_measurable_uminus_extreal[OF this] show ?r by simp
  15.168 +  assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
  15.169  qed auto
  15.170  
  15.171 -lemma (in sigma_algebra) borel_measurable_eq_atMost_extreal:
  15.172 -  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
  15.173 +lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal:
  15.174 +  fixes f :: "'a \<Rightarrow> ereal"
  15.175 +  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
  15.176  proof (intro iffI allI)
  15.177    assume pos[rule_format]: "\<forall>a. f -` {..a} \<inter> space M \<in> sets M"
  15.178    show "f \<in> borel_measurable M"
  15.179 -    unfolding borel_measurable_extreal_iff_real borel_measurable_iff_le
  15.180 +    unfolding borel_measurable_ereal_iff_real borel_measurable_iff_le
  15.181    proof (intro conjI allI)
  15.182      fix a :: real
  15.183 -    { fix x :: extreal assume *: "\<forall>i::nat. real i < x"
  15.184 +    { fix x :: ereal assume *: "\<forall>i::nat. real i < x"
  15.185        have "x = \<infinity>"
  15.186 -      proof (rule extreal_top)
  15.187 +      proof (rule ereal_top)
  15.188          fix B from real_arch_lt[of B] guess n ..
  15.189 -        then have "extreal B < real n" by auto
  15.190 +        then have "ereal B < real n" by auto
  15.191          with * show "B \<le> x" by (metis less_trans less_imp_le)
  15.192        qed }
  15.193      then have "f -` {\<infinity>} \<inter> space M = space M - (\<Union>i::nat. f -` {.. real i} \<inter> space M)"
  15.194        by (auto simp: not_le)
  15.195      then show "f -` {\<infinity>} \<inter> space M \<in> sets M" using pos by (auto simp del: UN_simps intro!: Diff)
  15.196      moreover
  15.197 -    have "{-\<infinity>} = {..-\<infinity>}" by auto
  15.198 +    have "{-\<infinity>::ereal} = {..-\<infinity>}" by auto
  15.199      then show "f -` {-\<infinity>} \<inter> space M \<in> sets M" using pos by auto
  15.200 -    moreover have "{x\<in>space M. f x \<le> extreal a} \<in> sets M"
  15.201 -      using pos[of "extreal a"] by (simp add: vimage_def Int_def conj_commute)
  15.202 +    moreover have "{x\<in>space M. f x \<le> ereal a} \<in> sets M"
  15.203 +      using pos[of "ereal a"] by (simp add: vimage_def Int_def conj_commute)
  15.204      moreover have "{w \<in> space M. real (f w) \<le> a} =
  15.205 -      (if a < 0 then {w \<in> space M. f w \<le> extreal a} - f -` {-\<infinity>} \<inter> space M
  15.206 -      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")
  15.207 +      (if a < 0 then {w \<in> space M. f w \<le> ereal a} - f -` {-\<infinity>} \<inter> space M
  15.208 +      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")
  15.209        proof (intro set_eqI) fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases "f x") auto qed
  15.210      ultimately show "{w \<in> space M. real (f w) \<le> a} \<in> sets M" by auto
  15.211    qed
  15.212  qed (simp add: measurable_sets)
  15.213  
  15.214 -lemma (in sigma_algebra) borel_measurable_eq_atLeast_extreal:
  15.215 -  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
  15.216 +lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal:
  15.217 +  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
  15.218  proof
  15.219    assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
  15.220    moreover have "\<And>a. (\<lambda>x. - f x) -` {..a} = f -` {-a ..}"
  15.221 -    by (auto simp: extreal_uminus_le_reorder)
  15.222 +    by (auto simp: ereal_uminus_le_reorder)
  15.223    ultimately have "(\<lambda>x. - f x) \<in> borel_measurable M"
  15.224 -    unfolding borel_measurable_eq_atMost_extreal by auto
  15.225 +    unfolding borel_measurable_eq_atMost_ereal by auto
  15.226    then show "f \<in> borel_measurable M" by simp
  15.227  qed (simp add: measurable_sets)
  15.228  
  15.229 -lemma (in sigma_algebra) borel_measurable_extreal_iff_less:
  15.230 -  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
  15.231 -  unfolding borel_measurable_eq_atLeast_extreal greater_eq_le_measurable ..
  15.232 +lemma (in sigma_algebra) borel_measurable_ereal_iff_less:
  15.233 +  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
  15.234 +  unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
  15.235  
  15.236 -lemma (in sigma_algebra) borel_measurable_extreal_iff_ge:
  15.237 -  "(f::'a \<Rightarrow> extreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
  15.238 -  unfolding borel_measurable_eq_atMost_extreal less_eq_ge_measurable ..
  15.239 +lemma (in sigma_algebra) borel_measurable_ereal_iff_ge:
  15.240 +  "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
  15.241 +  unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
  15.242  
  15.243 -lemma (in sigma_algebra) borel_measurable_extreal_eq_const:
  15.244 -  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
  15.245 +lemma (in sigma_algebra) borel_measurable_ereal_eq_const:
  15.246 +  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
  15.247    shows "{x\<in>space M. f x = c} \<in> sets M"
  15.248  proof -
  15.249    have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
  15.250    then show ?thesis using assms by (auto intro!: measurable_sets)
  15.251  qed
  15.252  
  15.253 -lemma (in sigma_algebra) borel_measurable_extreal_neq_const:
  15.254 -  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
  15.255 +lemma (in sigma_algebra) borel_measurable_ereal_neq_const:
  15.256 +  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
  15.257    shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
  15.258  proof -
  15.259    have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
  15.260    then show ?thesis using assms by (auto intro!: measurable_sets)
  15.261  qed
  15.262  
  15.263 -lemma (in sigma_algebra) borel_measurable_extreal_le[intro,simp]:
  15.264 -  fixes f g :: "'a \<Rightarrow> extreal"
  15.265 +lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]:
  15.266 +  fixes f g :: "'a \<Rightarrow> ereal"
  15.267    assumes f: "f \<in> borel_measurable M"
  15.268    assumes g: "g \<in> borel_measurable M"
  15.269    shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
  15.270 @@ -1283,13 +1285,13 @@
  15.271      {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>
  15.272      f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
  15.273    proof (intro set_eqI)
  15.274 -    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: extreal2_cases[of "f x" "g x"]) auto
  15.275 +    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: ereal2_cases[of "f x" "g x"]) auto
  15.276    qed
  15.277    with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
  15.278  qed
  15.279  
  15.280 -lemma (in sigma_algebra) borel_measurable_extreal_less[intro,simp]:
  15.281 -  fixes f :: "'a \<Rightarrow> extreal"
  15.282 +lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]:
  15.283 +  fixes f :: "'a \<Rightarrow> ereal"
  15.284    assumes f: "f \<in> borel_measurable M"
  15.285    assumes g: "g \<in> borel_measurable M"
  15.286    shows "{x \<in> space M. f x < g x} \<in> sets M"
  15.287 @@ -1298,8 +1300,8 @@
  15.288    then show ?thesis using g f by auto
  15.289  qed
  15.290  
  15.291 -lemma (in sigma_algebra) borel_measurable_extreal_eq[intro,simp]:
  15.292 -  fixes f :: "'a \<Rightarrow> extreal"
  15.293 +lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]:
  15.294 +  fixes f :: "'a \<Rightarrow> ereal"
  15.295    assumes f: "f \<in> borel_measurable M"
  15.296    assumes g: "g \<in> borel_measurable M"
  15.297    shows "{w \<in> space M. f w = g w} \<in> sets M"
  15.298 @@ -1308,8 +1310,8 @@
  15.299    then show ?thesis using g f by auto
  15.300  qed
  15.301  
  15.302 -lemma (in sigma_algebra) borel_measurable_extreal_neq[intro,simp]:
  15.303 -  fixes f :: "'a \<Rightarrow> extreal"
  15.304 +lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]:
  15.305 +  fixes f :: "'a \<Rightarrow> ereal"
  15.306    assumes f: "f \<in> borel_measurable M"
  15.307    assumes g: "g \<in> borel_measurable M"
  15.308    shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
  15.309 @@ -1323,23 +1325,23 @@
  15.310    "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
  15.311    by auto
  15.312  
  15.313 -lemma (in sigma_algebra) borel_measurable_extreal_add[intro, simp]:
  15.314 -  fixes f :: "'a \<Rightarrow> extreal"
  15.315 +lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]:
  15.316 +  fixes f :: "'a \<Rightarrow> ereal"
  15.317    assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
  15.318    shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
  15.319  proof -
  15.320    { fix x assume "x \<in> space M" then have "f x + g x =
  15.321        (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
  15.322          else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>
  15.323 -        else extreal (real (f x) + real (g x)))"
  15.324 -      by (cases rule: extreal2_cases[of "f x" "g x"]) auto }
  15.325 +        else ereal (real (f x) + real (g x)))"
  15.326 +      by (cases rule: ereal2_cases[of "f x" "g x"]) auto }
  15.327    with assms show ?thesis
  15.328      by (auto cong: measurable_cong simp: split_sets
  15.329               intro!: Un measurable_If measurable_sets)
  15.330  qed
  15.331  
  15.332 -lemma (in sigma_algebra) borel_measurable_extreal_setsum[simp, intro]:
  15.333 -  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
  15.334 +lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]:
  15.335 +  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
  15.336    assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
  15.337    shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
  15.338  proof cases
  15.339 @@ -1348,25 +1350,25 @@
  15.340      by induct auto
  15.341  qed (simp add: borel_measurable_const)
  15.342  
  15.343 -lemma (in sigma_algebra) borel_measurable_extreal_abs[intro, simp]:
  15.344 -  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M"
  15.345 +lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]:
  15.346 +  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
  15.347    shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
  15.348  proof -
  15.349    { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
  15.350    then show ?thesis using assms by (auto intro!: measurable_If)
  15.351  qed
  15.352  
  15.353 -lemma (in sigma_algebra) borel_measurable_extreal_times[intro, simp]:
  15.354 -  fixes f :: "'a \<Rightarrow> extreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
  15.355 +lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]:
  15.356 +  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
  15.357    shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
  15.358  proof -
  15.359 -  { fix f g :: "'a \<Rightarrow> extreal"
  15.360 +  { fix f g :: "'a \<Rightarrow> ereal"
  15.361      assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
  15.362        and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x"
  15.363      { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
  15.364          else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
  15.365 -        else extreal (real (f x) * real (g x)))"
  15.366 -      apply (cases rule: extreal2_cases[of "f x" "g x"])
  15.367 +        else ereal (real (f x) * real (g x)))"
  15.368 +      apply (cases rule: ereal2_cases[of "f x" "g x"])
  15.369        using pos[of x] by auto }
  15.370      with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M"
  15.371        by (auto cong: measurable_cong simp: split_sets
  15.372 @@ -1376,12 +1378,12 @@
  15.373      (\<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>))"
  15.374      by (auto simp: fun_eq_iff)
  15.375    show ?thesis using assms unfolding *
  15.376 -    by (intro measurable_If pos_times borel_measurable_uminus_extreal)
  15.377 +    by (intro measurable_If pos_times borel_measurable_uminus_ereal)
  15.378         (auto simp: split_sets intro!: Int)
  15.379  qed
  15.380  
  15.381 -lemma (in sigma_algebra) borel_measurable_extreal_setprod[simp, intro]:
  15.382 -  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> extreal"
  15.383 +lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]:
  15.384 +  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
  15.385    assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
  15.386    shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
  15.387  proof cases
  15.388 @@ -1389,25 +1391,25 @@
  15.389    thus ?thesis using assms by induct auto
  15.390  qed simp
  15.391  
  15.392 -lemma (in sigma_algebra) borel_measurable_extreal_min[simp, intro]:
  15.393 -  fixes f g :: "'a \<Rightarrow> extreal"
  15.394 +lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]:
  15.395 +  fixes f g :: "'a \<Rightarrow> ereal"
  15.396    assumes "f \<in> borel_measurable M"
  15.397    assumes "g \<in> borel_measurable M"
  15.398    shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
  15.399    using assms unfolding min_def by (auto intro!: measurable_If)
  15.400  
  15.401 -lemma (in sigma_algebra) borel_measurable_extreal_max[simp, intro]:
  15.402 -  fixes f g :: "'a \<Rightarrow> extreal"
  15.403 +lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]:
  15.404 +  fixes f g :: "'a \<Rightarrow> ereal"
  15.405    assumes "f \<in> borel_measurable M"
  15.406    and "g \<in> borel_measurable M"
  15.407    shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
  15.408    using assms unfolding max_def by (auto intro!: measurable_If)
  15.409  
  15.410  lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
  15.411 -  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> extreal"
  15.412 +  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
  15.413    assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
  15.414    shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
  15.415 -  unfolding borel_measurable_extreal_iff_ge
  15.416 +  unfolding borel_measurable_ereal_iff_ge
  15.417  proof
  15.418    fix a
  15.419    have "?sup -` {a<..} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
  15.420 @@ -1417,10 +1419,10 @@
  15.421  qed
  15.422  
  15.423  lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
  15.424 -  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> extreal"
  15.425 +  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"
  15.426    assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
  15.427    shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
  15.428 -  unfolding borel_measurable_extreal_iff_less
  15.429 +  unfolding borel_measurable_ereal_iff_less
  15.430  proof
  15.431    fix a
  15.432    have "?inf -` {..<a} \<inter> space M = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
  15.433 @@ -1430,30 +1432,30 @@
  15.434  qed
  15.435  
  15.436  lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:
  15.437 -  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
  15.438 +  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
  15.439    assumes "\<And>i. f i \<in> borel_measurable M"
  15.440    shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
  15.441    unfolding liminf_SUPR_INFI using assms by auto
  15.442  
  15.443  lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:
  15.444 -  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
  15.445 +  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
  15.446    assumes "\<And>i. f i \<in> borel_measurable M"
  15.447    shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
  15.448    unfolding limsup_INFI_SUPR using assms by auto
  15.449  
  15.450 -lemma (in sigma_algebra) borel_measurable_extreal_diff[simp, intro]:
  15.451 -  fixes f g :: "'a \<Rightarrow> extreal"
  15.452 +lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]:
  15.453 +  fixes f g :: "'a \<Rightarrow> ereal"
  15.454    assumes "f \<in> borel_measurable M"
  15.455    assumes "g \<in> borel_measurable M"
  15.456    shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
  15.457 -  unfolding minus_extreal_def using assms by auto
  15.458 +  unfolding minus_ereal_def using assms by auto
  15.459  
  15.460  lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:
  15.461 -  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
  15.462 +  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
  15.463    assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
  15.464    shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
  15.465    apply (subst measurable_cong)
  15.466 -  apply (subst suminf_extreal_eq_SUPR)
  15.467 +  apply (subst suminf_ereal_eq_SUPR)
  15.468    apply (rule pos)
  15.469    using assms by auto
  15.470  
  15.471 @@ -1465,11 +1467,11 @@
  15.472    and u: "\<And>i. u i \<in> borel_measurable M"
  15.473    shows "u' \<in> borel_measurable M"
  15.474  proof -
  15.475 -  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. extreal (u n x)) = extreal (u' x)"
  15.476 -    using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_extreal)
  15.477 -  moreover from u have "(\<lambda>x. liminf (\<lambda>n. extreal (u n x))) \<in> borel_measurable M"
  15.478 +  have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
  15.479 +    using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_ereal)
  15.480 +  moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
  15.481      by auto
  15.482 -  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_extreal_iff)
  15.483 +  ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
  15.484  qed
  15.485  
  15.486  end
    16.1 --- a/src/HOL/Probability/Caratheodory.thy	Wed Jul 20 10:11:08 2011 +0200
    16.2 +++ b/src/HOL/Probability/Caratheodory.thy	Wed Jul 20 10:48:00 2011 +0200
    16.3 @@ -19,8 +19,8 @@
    16.4    Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
    16.5  *}
    16.6  
    16.7 -lemma suminf_extreal_2dimen:
    16.8 -  fixes f:: "nat \<times> nat \<Rightarrow> extreal"
    16.9 +lemma suminf_ereal_2dimen:
   16.10 +  fixes f:: "nat \<times> nat \<Rightarrow> ereal"
   16.11    assumes pos: "\<And>p. 0 \<le> f p"
   16.12    assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
   16.13    shows "(\<Sum>i. f (prod_decode i)) = suminf g"
   16.14 @@ -47,21 +47,21 @@
   16.15    ultimately
   16.16    show ?thesis unfolding g_def using pos
   16.17      by (auto intro!: SUPR_eq  simp: setsum_cartesian_product reindex le_SUPI2
   16.18 -                     setsum_nonneg suminf_extreal_eq_SUPR SUPR_pair
   16.19 -                     SUPR_extreal_setsum[symmetric] incseq_setsumI setsum_nonneg)
   16.20 +                     setsum_nonneg suminf_ereal_eq_SUPR SUPR_pair
   16.21 +                     SUPR_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
   16.22  qed
   16.23  
   16.24  subsection {* Measure Spaces *}
   16.25  
   16.26  record 'a measure_space = "'a algebra" +
   16.27 -  measure :: "'a set \<Rightarrow> extreal"
   16.28 +  measure :: "'a set \<Rightarrow> ereal"
   16.29  
   16.30 -definition positive where "positive M f \<longleftrightarrow> f {} = (0::extreal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
   16.31 +definition positive where "positive M f \<longleftrightarrow> f {} = (0::ereal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
   16.32  
   16.33  definition additive where "additive M f \<longleftrightarrow>
   16.34    (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
   16.35  
   16.36 -definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> extreal) \<Rightarrow> bool" where
   16.37 +definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
   16.38    "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>
   16.39      (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
   16.40  
   16.41 @@ -168,7 +168,7 @@
   16.42    by (simp add: lambda_system_def)
   16.43  
   16.44  lemma (in algebra) lambda_system_Compl:
   16.45 -  fixes f:: "'a set \<Rightarrow> extreal"
   16.46 +  fixes f:: "'a set \<Rightarrow> ereal"
   16.47    assumes x: "x \<in> lambda_system M f"
   16.48    shows "space M - x \<in> lambda_system M f"
   16.49  proof -
   16.50 @@ -181,7 +181,7 @@
   16.51  qed
   16.52  
   16.53  lemma (in algebra) lambda_system_Int:
   16.54 -  fixes f:: "'a set \<Rightarrow> extreal"
   16.55 +  fixes f:: "'a set \<Rightarrow> ereal"
   16.56    assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
   16.57    shows "x \<inter> y \<in> lambda_system M f"
   16.58  proof -
   16.59 @@ -215,7 +215,7 @@
   16.60  qed
   16.61  
   16.62  lemma (in algebra) lambda_system_Un:
   16.63 -  fixes f:: "'a set \<Rightarrow> extreal"
   16.64 +  fixes f:: "'a set \<Rightarrow> ereal"
   16.65    assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
   16.66    shows "x \<union> y \<in> lambda_system M f"
   16.67  proof -
   16.68 @@ -321,7 +321,7 @@
   16.69  qed
   16.70  
   16.71  lemma (in algebra) increasing_additive_bound:
   16.72 -  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> extreal"
   16.73 +  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
   16.74    assumes f: "positive M f" and ad: "additive M f"
   16.75        and inc: "increasing M f"
   16.76        and A: "range A \<subseteq> sets M"
   16.77 @@ -346,7 +346,7 @@
   16.78    by (simp add: positive_def lambda_system_def)
   16.79  
   16.80  lemma (in algebra) lambda_system_strong_sum:
   16.81 -  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> extreal"
   16.82 +  fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
   16.83    assumes f: "positive M f" and a: "a \<in> sets M"
   16.84        and A: "range A \<subseteq> lambda_system M f"
   16.85        and disj: "disjoint_family A"
   16.86 @@ -537,7 +537,7 @@
   16.87    assumes posf: "positive M f" and ca: "countably_additive M f"
   16.88        and s: "s \<in> sets M"
   16.89    shows "Inf (measure_set M f s) = f s"
   16.90 -  unfolding Inf_extreal_def
   16.91 +  unfolding Inf_ereal_def
   16.92  proof (safe intro!: Greatest_equality)
   16.93    fix z
   16.94    assume z: "z \<in> measure_set M f s"
   16.95 @@ -648,7 +648,7 @@
   16.96  qed
   16.97  
   16.98  lemma (in ring_of_sets) inf_measure_close:
   16.99 -  fixes e :: extreal
  16.100 +  fixes e :: ereal
  16.101    assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
  16.102    shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
  16.103                 (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
  16.104 @@ -656,7 +656,7 @@
  16.105    from `Inf (measure_set M f s) \<noteq> \<infinity>` have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
  16.106      using inf_measure_pos[OF posf, of s] by auto
  16.107    obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
  16.108 -    using Inf_extreal_close[OF fin e] by auto
  16.109 +    using Inf_ereal_close[OF fin e] by auto
  16.110    thus ?thesis
  16.111      by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
  16.112  qed
  16.113 @@ -672,11 +672,11 @@
  16.114       and disj: "disjoint_family A"
  16.115       and sb: "(\<Union>i. A i) \<subseteq> space M"
  16.116  
  16.117 -  { fix e :: extreal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
  16.118 +  { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
  16.119      hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
  16.120          A n \<subseteq> (\<Union>i. BB n i) \<and> (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
  16.121        apply (safe intro!: choice inf_measure_close [of f, OF posf])
  16.122 -      using e sb by (auto simp: extreal_zero_less_0_iff one_extreal_def)
  16.123 +      using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def)
  16.124      then obtain BB
  16.125        where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
  16.126        and disjBB: "\<And>n. disjoint_family (BB n)"
  16.127 @@ -686,15 +686,15 @@
  16.128      have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n)) + e"
  16.129      proof -
  16.130        have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e"
  16.131 -        using suminf_half_series_extreal e
  16.132 -        by (simp add: extreal_zero_le_0_iff zero_le_divide_extreal suminf_cmult_extreal)
  16.133 +        using suminf_half_series_ereal e
  16.134 +        by (simp add: ereal_zero_le_0_iff zero_le_divide_ereal suminf_cmult_ereal)
  16.135        have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto
  16.136        then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le)
  16.137        then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n) + e*(1/2) ^ Suc n)"
  16.138          by (rule suminf_le_pos[OF BBle])
  16.139        also have "... = (\<Sum>n. ?outer (A n)) + e"
  16.140          using sum_eq_1 inf_measure_pos[OF posf] e
  16.141 -        by (subst suminf_add_extreal) (auto simp add: extreal_zero_le_0_iff)
  16.142 +        by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff)
  16.143        finally show ?thesis .
  16.144      qed
  16.145      def C \<equiv> "(split BB) o prod_decode"
  16.146 @@ -716,7 +716,7 @@
  16.147        by (rule ext)  (auto simp add: C_def)
  16.148      moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle
  16.149        using BB posf[unfolded positive_def]
  16.150 -      by (force intro!: suminf_extreal_2dimen simp: o_def)
  16.151 +      by (force intro!: suminf_ereal_2dimen simp: o_def)
  16.152      ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def)
  16.153      have "?outer (\<Union>i. A i) \<le> (\<Sum>n. \<Sum>i. f (BB n i))"
  16.154        apply (rule inf_measure_le [OF posf(1) inc], auto)
  16.155 @@ -732,7 +732,7 @@
  16.156    proof cases
  16.157      assume "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
  16.158      with for_finite_Inf show ?thesis
  16.159 -      by (intro extreal_le_epsilon) auto
  16.160 +      by (intro ereal_le_epsilon) auto
  16.161    next
  16.162      assume "\<not> (\<forall>i. ?outer (A i) \<noteq> \<infinity>)"
  16.163      then have "\<exists>i. ?outer (A i) = \<infinity>"
  16.164 @@ -771,7 +771,7 @@
  16.165    next
  16.166      assume fin: "Inf (measure_set M f s) \<noteq> \<infinity>"
  16.167      then have "measure_set M f s \<noteq> {}"
  16.168 -      by (auto simp: top_extreal_def)
  16.169 +      by (auto simp: top_ereal_def)
  16.170      show ?thesis
  16.171      proof (rule complete_lattice_class.Inf_greatest)
  16.172        fix r assume "r \<in> measure_set M f s"
  16.173 @@ -793,7 +793,7 @@
  16.174        ultimately have "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le>
  16.175            (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
  16.176        also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
  16.177 -        using A(2) x posf by (subst suminf_add_extreal) (auto simp: positive_def)
  16.178 +        using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def)
  16.179        also have "\<dots> = (\<Sum>i. f (A i))"
  16.180          using A x
  16.181          by (subst add[THEN additiveD, symmetric])
  16.182 @@ -830,7 +830,7 @@
  16.183  
  16.184  theorem (in ring_of_sets) caratheodory:
  16.185    assumes posf: "positive M f" and ca: "countably_additive M f"
  16.186 -  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
  16.187 +  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
  16.188              measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
  16.189  proof -
  16.190    have inc: "increasing M f"
  16.191 @@ -873,7 +873,7 @@
  16.192      by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
  16.193    moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
  16.194      using f(1)[unfolded positive_def] dA
  16.195 -    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_extreal_pos)
  16.196 +    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
  16.197    from LIMSEQ_Suc[OF this]
  16.198    have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
  16.199      unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
  16.200 @@ -936,13 +936,13 @@
  16.201      show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> sets M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
  16.202        using A by auto
  16.203    qed
  16.204 -  from INF_Lim_extreal[OF decseq_f this]
  16.205 +  from INF_Lim_ereal[OF decseq_f this]
  16.206    have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
  16.207    moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
  16.208      by auto
  16.209    ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
  16.210      using A(4) f_fin f_Int_fin
  16.211 -    by (subst INFI_extreal_add) (auto simp: decseq_f)
  16.212 +    by (subst INFI_ereal_add) (auto simp: decseq_f)
  16.213    moreover {
  16.214      fix n
  16.215      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))"
  16.216 @@ -952,7 +952,7 @@
  16.217      finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
  16.218    ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
  16.219      by simp
  16.220 -  with LIMSEQ_extreal_INFI[OF decseq_fA]
  16.221 +  with LIMSEQ_ereal_INFI[OF decseq_fA]
  16.222    show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
  16.223  qed
  16.224  
  16.225 @@ -965,9 +965,9 @@
  16.226    assumes A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
  16.227    shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
  16.228  proof -
  16.229 -  have "\<forall>A\<in>sets M. \<exists>x. f A = extreal x"
  16.230 +  have "\<forall>A\<in>sets M. \<exists>x. f A = ereal x"
  16.231    proof
  16.232 -    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = extreal x"
  16.233 +    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = ereal x"
  16.234        unfolding positive_def by (cases "f A") auto
  16.235    qed
  16.236    from bchoice[OF this] guess f' .. note f' = this[rule_format]
  16.237 @@ -981,10 +981,10 @@
  16.238        by auto
  16.239      finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)"
  16.240        using A by (subst (asm) (1 2 3) f') auto
  16.241 -    then have "f ((\<Union>i. A i) - A i) = extreal (f' (\<Union>i. A i) - f' (A i))"
  16.242 +    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
  16.243        using A f' by auto }
  16.244    ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
  16.245 -    by (simp add: zero_extreal_def)
  16.246 +    by (simp add: zero_ereal_def)
  16.247    then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
  16.248      by (rule LIMSEQ_diff_approach_zero2[OF LIMSEQ_const])
  16.249    then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
  16.250 @@ -1002,7 +1002,7 @@
  16.251  lemma (in ring_of_sets) caratheodory_empty_continuous:
  16.252    assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> sets M \<Longrightarrow> f A \<noteq> \<infinity>"
  16.253    assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
  16.254 -  shows "\<exists>\<mu> :: 'a set \<Rightarrow> extreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
  16.255 +  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
  16.256              measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
  16.257  proof (intro caratheodory empty_continuous_imp_countably_additive f)
  16.258    show "\<forall>A\<in>sets M. f A \<noteq> \<infinity>" using fin by auto
    17.1 --- a/src/HOL/Probability/Complete_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    17.2 +++ b/src/HOL/Probability/Complete_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    17.3 @@ -253,7 +253,7 @@
    17.4  qed
    17.5  
    17.6  lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
    17.7 -  fixes g :: "'a \<Rightarrow> extreal"
    17.8 +  fixes g :: "'a \<Rightarrow> ereal"
    17.9    assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
   17.10    shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
   17.11  proof -
   17.12 @@ -279,7 +279,7 @@
   17.13  qed
   17.14  
   17.15  lemma (in completeable_measure_space) completion_ex_borel_measurable:
   17.16 -  fixes g :: "'a \<Rightarrow> extreal"
   17.17 +  fixes g :: "'a \<Rightarrow> ereal"
   17.18    assumes g: "g \<in> borel_measurable completion"
   17.19    shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
   17.20  proof -
    18.1 --- a/src/HOL/Probability/Conditional_Probability.thy	Wed Jul 20 10:11:08 2011 +0200
    18.2 +++ b/src/HOL/Probability/Conditional_Probability.thy	Wed Jul 20 10:48:00 2011 +0200
    18.3 @@ -15,7 +15,7 @@
    18.4      \<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)))"
    18.5  
    18.6  lemma (in prob_space) conditional_expectation_exists:
    18.7 -  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
    18.8 +  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
    18.9    assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   18.10    and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
   18.11    shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
   18.12 @@ -52,7 +52,7 @@
   18.13  qed
   18.14  
   18.15  lemma (in prob_space)
   18.16 -  fixes X :: "'a \<Rightarrow> extreal" and N :: "('a, 'b) measure_space_scheme"
   18.17 +  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
   18.18    assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
   18.19    and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
   18.20    shows borel_measurable_conditional_expectation:
   18.21 @@ -71,7 +71,7 @@
   18.22  qed
   18.23  
   18.24  lemma (in sigma_algebra) factorize_measurable_function_pos:
   18.25 -  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
   18.26 +  fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
   18.27    assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
   18.28    assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
   18.29    shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
   18.30 @@ -98,7 +98,7 @@
   18.31        show "simple_function M' ?g" using B by auto
   18.32  
   18.33        fix x assume "x \<in> space M"
   18.34 -      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)"
   18.35 +      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)"
   18.36          unfolding indicator_def using B by auto
   18.37        then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
   18.38          by (subst va.simple_function_indicator_representation) auto
   18.39 @@ -119,7 +119,7 @@
   18.40  qed
   18.41  
   18.42  lemma (in sigma_algebra) factorize_measurable_function:
   18.43 -  fixes Z :: "'a \<Rightarrow> extreal" and Y :: "'a \<Rightarrow> 'c"
   18.44 +  fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
   18.45    assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
   18.46    shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
   18.47      \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
   18.48 @@ -129,7 +129,7 @@
   18.49    from M'.sigma_algebra_vimage[OF this]
   18.50    interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
   18.51  
   18.52 -  { fix g :: "'c \<Rightarrow> extreal" assume "g \<in> borel_measurable M'"
   18.53 +  { fix g :: "'c \<Rightarrow> ereal" assume "g \<in> borel_measurable M'"
   18.54      with M'.measurable_vimage_algebra[OF Y]
   18.55      have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
   18.56        by (rule measurable_comp)
    19.1 --- a/src/HOL/Probability/Finite_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    19.2 +++ b/src/HOL/Probability/Finite_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    19.3 @@ -538,13 +538,13 @@
    19.4  next
    19.5    assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
    19.6    then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
    19.7 -    by (auto simp: setprod_extreal_0 intro!: bexI)
    19.8 +    by (auto simp: setprod_ereal_0 intro!: bexI)
    19.9    moreover
   19.10    have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
   19.11      by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
   19.12         (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
   19.13    then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
   19.14 -    by (auto simp: setprod_extreal_0 intro!: bexI)
   19.15 +    by (auto simp: setprod_ereal_0 intro!: bexI)
   19.16    ultimately show ?thesis
   19.17      unfolding product_algebra_generator_def by simp
   19.18  qed
   19.19 @@ -601,7 +601,7 @@
   19.20  using `finite I` proof induct
   19.21    case empty
   19.22    interpret finite_product_sigma_finite M "{}" by default simp
   19.23 -  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> extreal"
   19.24 +  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> ereal"
   19.25    show ?case
   19.26    proof (intro exI conjI ballI)
   19.27      have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
   19.28 @@ -858,7 +858,7 @@
   19.29  qed
   19.30  
   19.31  lemma (in product_sigma_finite) product_positive_integral_setprod:
   19.32 -  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> extreal"
   19.33 +  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
   19.34    assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
   19.35    and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
   19.36    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))"
   19.37 @@ -874,14 +874,14 @@
   19.38      using insert by (auto intro!: setprod_cong)
   19.39    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)"
   19.40      using sets_into_space insert
   19.41 -    by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra
   19.42 +    by (intro sigma_algebra.borel_measurable_ereal_setprod sigma_algebra_product_algebra
   19.43                measurable_comp[OF measurable_component_singleton, unfolded comp_def])
   19.44         auto
   19.45    then show ?case
   19.46      apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
   19.47 -    apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc)
   19.48 +    apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc)
   19.49      apply (subst I.positive_integral_cmult)
   19.50 -    apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive)
   19.51 +    apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive)
   19.52      done
   19.53  qed
   19.54  
   19.55 @@ -890,8 +890,8 @@
   19.56    shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
   19.57  proof -
   19.58    interpret I: finite_product_sigma_finite M "{i}" by default simp
   19.59 -  have *: "(\<lambda>x. extreal (f x)) \<in> borel_measurable (M i)"
   19.60 -    "(\<lambda>x. extreal (- f x)) \<in> borel_measurable (M i)"
   19.61 +  have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)"
   19.62 +    "(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)"
   19.63      using assms by auto
   19.64    show ?thesis
   19.65      unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
   19.66 @@ -965,17 +965,17 @@
   19.67    proof (unfold integrable_def, intro conjI)
   19.68      show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
   19.69        using borel by auto
   19.70 -    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)"
   19.71 -      by (simp add: setprod_extreal abs_setprod)
   19.72 -    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. extreal (abs (f i x)) \<partial>M i))"
   19.73 +    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)"
   19.74 +      by (simp add: setprod_ereal abs_setprod)
   19.75 +    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))"
   19.76        using f by (subst product_positive_integral_setprod) auto
   19.77      also have "\<dots> < \<infinity>"
   19.78        using integrable[THEN M.integrable_abs]
   19.79        by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
   19.80 -    finally show "(\<integral>\<^isup>+x. extreal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
   19.81 -    have "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
   19.82 +    finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
   19.83 +    have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
   19.84        by (intro positive_integral_cong_pos) auto
   19.85 -    then show "(\<integral>\<^isup>+x. extreal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
   19.86 +    then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
   19.87    qed
   19.88    ultimately show ?thesis
   19.89      by (rule integrable_abs_iff[THEN iffD1])
    20.1 --- a/src/HOL/Probability/Independent_Family.thy	Wed Jul 20 10:11:08 2011 +0200
    20.2 +++ b/src/HOL/Probability/Independent_Family.thy	Wed Jul 20 10:48:00 2011 +0200
    20.3 @@ -822,9 +822,9 @@
    20.4    assumes I: "I \<noteq> {}" "finite I"
    20.5    assumes rv: "\<And>i. random_variable (M' i) (X i)"
    20.6    shows "indep_vars M' X I \<longleftrightarrow>
    20.7 -    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := extreal\<circ>distribution (X i) \<rparr>)).
    20.8 +    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)).
    20.9        distribution (\<lambda>x. \<lambda>i\<in>I. X i x) A =
   20.10 -      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := extreal\<circ>distribution (X i) \<rparr>)) A)"
   20.11 +      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)) A)"
   20.12      (is "_ \<longleftrightarrow> (\<forall>X\<in>_. distribution ?D X = finite_measure.\<mu>' (Pi\<^isub>M I ?M) X)")
   20.13  proof -
   20.14    interpret M': prob_space "?M i" for i
   20.15 @@ -832,7 +832,7 @@
   20.16    interpret P: finite_product_prob_space ?M I
   20.17      proof qed fact
   20.18  
   20.19 -  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := extreal \<circ> distribution ?D \<rparr>"
   20.20 +  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := ereal \<circ> distribution ?D \<rparr>"
   20.21    have "random_variable P.P ?D"
   20.22      using `finite I` rv by (intro random_variable_restrict) auto
   20.23    then interpret D: prob_space ?D'
   20.24 @@ -938,24 +938,24 @@
   20.25  
   20.26  lemma (in prob_space) indep_var_distributionD:
   20.27    assumes indep: "indep_var S X T Y"
   20.28 -  defines "P \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
   20.29 +  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   20.30    assumes "A \<in> sets P"
   20.31    shows "joint_distribution X Y A = finite_measure.\<mu>' P A"
   20.32  proof -
   20.33    from indep have rvs: "random_variable S X" "random_variable T Y"
   20.34      by (blast dest: indep_var_rv1 indep_var_rv2)+
   20.35  
   20.36 -  let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
   20.37 -  let ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
   20.38 +  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
   20.39 +  let ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   20.40    interpret X: prob_space ?S by (rule distribution_prob_space) fact
   20.41    interpret Y: prob_space ?T by (rule distribution_prob_space) fact
   20.42    interpret XY: pair_prob_space ?S ?T by default
   20.43  
   20.44 -  let ?J = "XY.P\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>"
   20.45 +  let ?J = "XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>"
   20.46    interpret J: prob_space ?J
   20.47      by (rule joint_distribution_prob_space) (simp_all add: rvs)
   20.48  
   20.49 -  have "finite_measure.\<mu>' (XY.P\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
   20.50 +  have "finite_measure.\<mu>' (XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
   20.51    proof (rule prob_space_unique_Int_stable)
   20.52      show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P")
   20.53        by fact
    21.1 --- a/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    21.2 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    21.3 @@ -563,7 +563,7 @@
    21.4              using `0 \<le> ?a` Q_sets J'.measure_space_1
    21.5              by (subst J'.positive_integral_add) auto
    21.6            finally show "?a / 2^(k+1) \<le> measure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
    21.7 -            by (cases rule: extreal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
    21.8 +            by (cases rule: ereal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
    21.9                 (auto simp: field_simps)
   21.10          qed
   21.11          also have "\<dots> = measure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
   21.12 @@ -712,7 +712,7 @@
   21.13        with `(\<Inter>i. A i) = {}` show False by auto
   21.14      qed
   21.15      ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
   21.16 -      using LIMSEQ_extreal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
   21.17 +      using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
   21.18    qed
   21.19  qed
   21.20  
   21.21 @@ -812,7 +812,7 @@
   21.22    using assms
   21.23    unfolding \<mu>'_def M.\<mu>'_def
   21.24    by (subst measure_times[OF assms])
   21.25 -     (auto simp: finite_measure_eq M.finite_measure_eq setprod_extreal)
   21.26 +     (auto simp: finite_measure_eq M.finite_measure_eq setprod_ereal)
   21.27  
   21.28  lemma (in product_prob_space) finite_measure_infprod_emb_Pi:
   21.29    assumes J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> X j \<in> sets (M j)"
    22.1 --- a/src/HOL/Probability/Information.thy	Wed Jul 20 10:11:08 2011 +0200
    22.2 +++ b/src/HOL/Probability/Information.thy	Wed Jul 20 10:48:00 2011 +0200
    22.3 @@ -203,7 +203,7 @@
    22.4  
    22.5    from real_RN_deriv[OF fms ac] guess D . note D = this
    22.6    with absolutely_continuous_AE[OF ms] ac
    22.7 -  have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = extreal (D x)"
    22.8 +  have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = ereal (D x)"
    22.9      by auto
   22.10  
   22.11    def f \<equiv> "\<lambda>x. if D x = 0 then 1 else 1 / D x"
   22.12 @@ -215,7 +215,7 @@
   22.13      by (simp add: integral_cmult)
   22.14  
   22.15    { fix A assume "A \<in> sets M"
   22.16 -    with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. extreal (D x) * indicator A x \<partial>M)"
   22.17 +    with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. ereal (D x) * indicator A x \<partial>M)"
   22.18        by (auto intro!: positive_integral_cong_AE) }
   22.19    note D_density = this
   22.20  
   22.21 @@ -231,12 +231,12 @@
   22.22    ultimately have M_int: "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
   22.23      by simp
   22.24  
   22.25 -  have D_neg: "(\<integral>\<^isup>+ x. extreal (- D x) \<partial>M) = 0"
   22.26 +  have D_neg: "(\<integral>\<^isup>+ x. ereal (- D x) \<partial>M) = 0"
   22.27      using D by (subst positive_integral_0_iff_AE) auto
   22.28  
   22.29 -  have "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = \<nu> (space M)"
   22.30 +  have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = \<nu> (space M)"
   22.31      using RN D by (auto intro!: positive_integral_cong_AE)
   22.32 -  then have D_pos: "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = 1"
   22.33 +  then have D_pos: "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = 1"
   22.34      using \<nu>.measure_space_1 by simp
   22.35  
   22.36    have "integrable M D"
   22.37 @@ -271,16 +271,16 @@
   22.38  
   22.39        have "\<mu> {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
   22.40          using D(1) by auto
   22.41 -      also have "\<dots> = (\<integral>\<^isup>+ x. extreal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
   22.42 -        using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_extreal_def)
   22.43 +      also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
   22.44 +        using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def)
   22.45        also have "\<dots> = \<nu> {x\<in>space M. D x \<noteq> 0}"
   22.46          using D(1) D_density by auto
   22.47        also have "\<dots> = \<nu> (space M)"
   22.48          using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def)
   22.49        finally have "AE x. D x = 1"
   22.50          using D(1) \<nu>.measure_space_1 by (intro AE_I_eq_1) auto
   22.51 -      then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. extreal (D x) * indicator A x\<partial>M)"
   22.52 -        by (intro positive_integral_cong_AE) (auto simp: one_extreal_def[symmetric])
   22.53 +      then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. ereal (D x) * indicator A x\<partial>M)"
   22.54 +        by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
   22.55        also have "\<dots> = \<nu> A"
   22.56          using `A \<in> sets M` D_density by simp
   22.57        finally show False using `A \<in> sets M` `\<nu> A \<noteq> \<mu> A` by simp
   22.58 @@ -293,7 +293,7 @@
   22.59        using D(2)
   22.60      proof (elim AE_mp, safe intro!: AE_I2)
   22.61        fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0"
   22.62 -        and RN: "RN_deriv M \<nu> t = extreal (D t)"
   22.63 +        and RN: "RN_deriv M \<nu> t = ereal (D t)"
   22.64          and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \<nu> t)"
   22.65  
   22.66        have "D t - 1 = D t - indicator ?D_set t"
   22.67 @@ -311,7 +311,7 @@
   22.68      show "AE t. D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
   22.69        using D(2)
   22.70      proof (elim AE_mp, intro AE_I2 impI)
   22.71 -      fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = extreal (D t)"
   22.72 +      fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = ereal (D t)"
   22.73        show "D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
   22.74        proof cases
   22.75          assume asm: "D t \<noteq> 0"
   22.76 @@ -347,7 +347,7 @@
   22.77        using eq by (intro measure_space_cong) auto
   22.78      show "absolutely_continuous \<nu>"
   22.79        unfolding absolutely_continuous_def using eq by auto
   22.80 -    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: extreal)" by auto
   22.81 +    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: ereal)" by auto
   22.82      fix A assume "A \<in> sets M"
   22.83      with eq show "\<nu> A = \<integral>\<^isup>+ x. 1 * indicator A x \<partial>M" by simp
   22.84    qed
   22.85 @@ -468,17 +468,17 @@
   22.86  
   22.87  definition (in prob_space)
   22.88    "mutual_information b S T X Y =
   22.89 -    KL_divergence b (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
   22.90 -      (extreal\<circ>joint_distribution X Y)"
   22.91 +    KL_divergence b (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
   22.92 +      (ereal\<circ>joint_distribution X Y)"
   22.93  
   22.94  lemma (in information_space)
   22.95    fixes S T X Y
   22.96 -  defines "P \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
   22.97 +  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
   22.98    shows "indep_var S X T Y \<longleftrightarrow>
   22.99      (random_variable S X \<and> random_variable T Y \<and>
  22.100 -      measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y) \<and>
  22.101 -      integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
  22.102 -        (entropy_density b P (extreal\<circ>joint_distribution X Y)) \<and>
  22.103 +      measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y) \<and>
  22.104 +      integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
  22.105 +        (entropy_density b P (ereal\<circ>joint_distribution X Y)) \<and>
  22.106       mutual_information b S T X Y = 0)"
  22.107  proof safe
  22.108    assume indep: "indep_var S X T Y"
  22.109 @@ -487,16 +487,16 @@
  22.110    then show "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
  22.111      by blast+
  22.112  
  22.113 -  interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
  22.114 +  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
  22.115      by (rule distribution_prob_space) fact
  22.116 -  interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
  22.117 +  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
  22.118      by (rule distribution_prob_space) fact
  22.119 -  interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
  22.120 +  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
  22.121    interpret XY: information_space XY.P b by default (rule b_gt_1)
  22.122  
  22.123 -  let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
  22.124 +  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
  22.125    { fix A assume "A \<in> sets XY.P"
  22.126 -    then have "extreal (joint_distribution X Y A) = XY.\<mu> A"
  22.127 +    then have "ereal (joint_distribution X Y A) = XY.\<mu> A"
  22.128        using indep_var_distributionD[OF indep]
  22.129        by (simp add: XY.P.finite_measure_eq) }
  22.130    note j_eq = this
  22.131 @@ -504,31 +504,31 @@
  22.132    interpret J: prob_space ?J
  22.133      using j_eq by (intro XY.prob_space_cong) auto
  22.134  
  22.135 -  have ac: "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
  22.136 +  have ac: "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
  22.137      by (simp add: XY.absolutely_continuous_def j_eq)
  22.138 -  then show "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
  22.139 +  then show "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
  22.140      unfolding P_def .
  22.141  
  22.142 -  have ed: "entropy_density b XY.P (extreal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
  22.143 +  have ed: "entropy_density b XY.P (ereal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
  22.144      by (rule XY.measurable_entropy_density) (default | fact)+
  22.145  
  22.146 -  have "AE x in XY.P. 1 = RN_deriv XY.P (extreal\<circ>joint_distribution X Y) x"
  22.147 +  have "AE x in XY.P. 1 = RN_deriv XY.P (ereal\<circ>joint_distribution X Y) x"
  22.148    proof (rule XY.RN_deriv_unique[OF _ ac])
  22.149      show "measure_space ?J" by default
  22.150      fix A assume "A \<in> sets XY.P"
  22.151 -    then show "(extreal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)"
  22.152 +    then show "(ereal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)"
  22.153        by (simp add: j_eq)
  22.154    qed (insert XY.measurable_const[of 1 borel], auto)
  22.155 -  then have ae_XY: "AE x in XY.P. entropy_density b XY.P (extreal\<circ>joint_distribution X Y) x = 0"
  22.156 +  then have ae_XY: "AE x in XY.P. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
  22.157      by (elim XY.AE_mp) (simp add: entropy_density_def)
  22.158 -  have ae_J: "AE x in ?J. entropy_density b XY.P (extreal\<circ>joint_distribution X Y) x = 0"
  22.159 +  have ae_J: "AE x in ?J. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
  22.160    proof (rule XY.absolutely_continuous_AE)
  22.161      show "measure_space ?J" by default
  22.162      show "XY.absolutely_continuous (measure ?J)"
  22.163        using ac by simp
  22.164    qed (insert ae_XY, simp_all)
  22.165 -  then show "integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
  22.166 -        (entropy_density b P (extreal\<circ>joint_distribution X Y))"
  22.167 +  then show "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
  22.168 +        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
  22.169      unfolding P_def
  22.170      using ed XY.measurable_const[of 0 borel]
  22.171      by (subst J.integrable_cong_AE) auto
  22.172 @@ -540,30 +540,30 @@
  22.173    assume "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
  22.174    then have rvs: "random_variable S X" "random_variable T Y" by blast+
  22.175  
  22.176 -  interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
  22.177 +  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
  22.178      by (rule distribution_prob_space) fact
  22.179 -  interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
  22.180 +  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
  22.181      by (rule distribution_prob_space) fact
  22.182 -  interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
  22.183 +  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
  22.184    interpret XY: information_space XY.P b by default (rule b_gt_1)
  22.185  
  22.186 -  let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
  22.187 +  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
  22.188    interpret J: prob_space ?J
  22.189      using rvs by (intro joint_distribution_prob_space) auto
  22.190  
  22.191 -  assume ac: "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
  22.192 -  assume int: "integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
  22.193 -        (entropy_density b P (extreal\<circ>joint_distribution X Y))"
  22.194 +  assume ac: "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
  22.195 +  assume int: "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
  22.196 +        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
  22.197    assume I_eq_0: "mutual_information b S T X Y = 0"
  22.198  
  22.199 -  have eq: "\<forall>A\<in>sets XY.P. (extreal \<circ> joint_distribution X Y) A = XY.\<mu> A"
  22.200 +  have eq: "\<forall>A\<in>sets XY.P. (ereal \<circ> joint_distribution X Y) A = XY.\<mu> A"
  22.201    proof (rule XY.KL_eq_0_imp)
  22.202      show "prob_space ?J" by default
  22.203 -    show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
  22.204 +    show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
  22.205        using ac by (simp add: P_def)
  22.206 -    show "integrable ?J (entropy_density b XY.P (extreal\<circ>joint_distribution X Y))"
  22.207 +    show "integrable ?J (entropy_density b XY.P (ereal\<circ>joint_distribution X Y))"
  22.208        using int by (simp add: P_def)
  22.209 -    show "KL_divergence b XY.P (extreal\<circ>joint_distribution X Y) = 0"
  22.210 +    show "KL_divergence b XY.P (ereal\<circ>joint_distribution X Y) = 0"
  22.211        using I_eq_0 unfolding mutual_information_def by (simp add: P_def)
  22.212    qed
  22.213  
  22.214 @@ -587,13 +587,13 @@
  22.215          using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
  22.216      next
  22.217        fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
  22.218 -      have "extreal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
  22.219 -        extreal (joint_distribution X Y (A \<times> B))"
  22.220 +      have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
  22.221 +        ereal (joint_distribution X Y (A \<times> B))"
  22.222          unfolding distribution_def
  22.223 -        by (intro arg_cong[where f="\<lambda>C. extreal (prob C)"]) auto
  22.224 +        by (intro arg_cong[where f="\<lambda>C. ereal (prob C)"]) auto
  22.225        also have "\<dots> = XY.\<mu> (A \<times> B)"
  22.226          using ab eq by (auto simp: XY.finite_measure_eq)
  22.227 -      also have "\<dots> = extreal (distribution X A) * extreal (distribution Y B)"
  22.228 +      also have "\<dots> = ereal (distribution X A) * ereal (distribution Y B)"
  22.229          using ab by (simp add: XY.pair_measure_times)
  22.230        finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
  22.231          prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
  22.232 @@ -605,10 +605,10 @@
  22.233  lemma (in information_space) mutual_information_commute_generic:
  22.234    assumes X: "random_variable S X" and Y: "random_variable T Y"
  22.235    assumes ac: "measure_space.absolutely_continuous
  22.236 -    (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)"
  22.237 +    (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)"
  22.238    shows "mutual_information b S T X Y = mutual_information b T S Y X"
  22.239  proof -
  22.240 -  let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
  22.241 +  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
  22.242    interpret S: prob_space ?S using X by (rule distribution_prob_space)
  22.243    interpret T: prob_space ?T using Y by (rule distribution_prob_space)
  22.244    interpret P: pair_prob_space ?S ?T ..
  22.245 @@ -617,13 +617,13 @@
  22.246      unfolding mutual_information_def
  22.247    proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1])
  22.248      show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
  22.249 -      (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
  22.250 +      (P.P \<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := ereal\<circ>joint_distribution Y X\<rparr>)"
  22.251        using X Y unfolding measurable_def
  22.252        unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable
  22.253        by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>'])
  22.254 -    have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
  22.255 +    have "prob_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
  22.256        using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
  22.257 -    then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
  22.258 +    then show "measure_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
  22.259        unfolding prob_space_def by simp
  22.260    qed auto
  22.261  qed
  22.262 @@ -634,33 +634,33 @@
  22.263  abbreviation (in information_space)
  22.264    mutual_information_Pow ("\<I>'(_ ; _')") where
  22.265    "\<I>(X ; Y) \<equiv> mutual_information b
  22.266 -    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
  22.267 -    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
  22.268 +    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
  22.269 +    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
  22.270  
  22.271  lemma (in prob_space) finite_variables_absolutely_continuous:
  22.272    assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
  22.273    shows "measure_space.absolutely_continuous
  22.274 -    (S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
  22.275 -    (extreal\<circ>joint_distribution X Y)"
  22.276 +    (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
  22.277 +    (ereal\<circ>joint_distribution X Y)"
  22.278  proof -
  22.279 -  interpret X: finite_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
  22.280 +  interpret X: finite_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
  22.281      using X by (rule distribution_finite_prob_space)
  22.282 -  interpret Y: finite_prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
  22.283 +  interpret Y: finite_prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
  22.284      using Y by (rule distribution_finite_prob_space)
  22.285    interpret XY: pair_finite_prob_space
  22.286 -    "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr> measure := extreal\<circ>distribution Y\<rparr>" by default
  22.287 -  interpret P: finite_prob_space "XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>"
  22.288 +    "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr> measure := ereal\<circ>distribution Y\<rparr>" by default
  22.289 +  interpret P: finite_prob_space "XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>"
  22.290      using assms by (auto intro!: joint_distribution_finite_prob_space)
  22.291    note rv = assms[THEN finite_random_variableD]
  22.292 -  show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
  22.293 +  show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
  22.294    proof (rule XY.absolutely_continuousI)
  22.295 -    show "finite_measure_space (XY.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
  22.296 +    show "finite_measure_space (XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
  22.297      fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
  22.298      then obtain a b where "x = (a, b)"
  22.299        and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
  22.300        by (cases x) (auto simp: space_pair_measure)
  22.301      with finite_distribution_order(5,6)[OF X Y]
  22.302 -    show "(extreal \<circ> joint_distribution X Y) {x} = 0" by auto
  22.303 +    show "(ereal \<circ> joint_distribution X Y) {x} = 0" by auto
  22.304    qed
  22.305  qed
  22.306  
  22.307 @@ -676,16 +676,16 @@
  22.308    and mutual_information_positive_generic:
  22.309       "0 \<le> mutual_information b MX MY X Y" (is ?positive)
  22.310  proof -
  22.311 -  interpret X: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
  22.312 +  interpret X: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
  22.313      using MX by (rule distribution_finite_prob_space)
  22.314 -  interpret Y: finite_prob_space "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
  22.315 +  interpret Y: finite_prob_space "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
  22.316      using MY by (rule distribution_finite_prob_space)
  22.317 -  interpret XY: pair_finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>" "MY\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
  22.318 -  interpret P: finite_prob_space "XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>"
  22.319 +  interpret XY: pair_finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
  22.320 +  interpret P: finite_prob_space "XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>"
  22.321      using assms by (auto intro!: joint_distribution_finite_prob_space)
  22.322  
  22.323 -  have P_ms: "finite_measure_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
  22.324 -  have P_ps: "finite_prob_space (XY.P\<lparr>measure := extreal\<circ>joint_distribution X Y\<rparr>)" by default
  22.325 +  have P_ms: "finite_measure_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
  22.326 +  have P_ps: "finite_prob_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by default
  22.327  
  22.328    show ?sum
  22.329      unfolding Let_def mutual_information_def
  22.330 @@ -739,14 +739,14 @@
  22.331  
  22.332  abbreviation (in information_space)
  22.333    entropy_Pow ("\<H>'(_')") where
  22.334 -  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr> X"
  22.335 +  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> X"
  22.336  
  22.337  lemma (in information_space) entropy_generic_eq:
  22.338    fixes X :: "'a \<Rightarrow> 'c"
  22.339    assumes MX: "finite_random_variable MX X"
  22.340    shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
  22.341  proof -
  22.342 -  interpret MX: finite_prob_space "MX\<lparr>measure := extreal\<circ>distribution X\<rparr>"
  22.343 +  interpret MX: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
  22.344      using MX by (rule distribution_finite_prob_space)
  22.345    let "?X x" = "distribution X {x}"
  22.346    let "?XX x y" = "joint_distribution X X {(x, y)}"
  22.347 @@ -779,9 +779,9 @@
  22.348    assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
  22.349    shows "\<H>(X) = 0"
  22.350  proof -
  22.351 -  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal\<circ>distribution X\<rparr>"
  22.352 +  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal\<circ>distribution X\<rparr>"
  22.353    note simple_function_imp_finite_random_variable[OF `simple_function M X`]
  22.354 -  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
  22.355 +  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
  22.356    interpret X: finite_prob_space ?X by simp
  22.357    have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
  22.358      using X.measure_compl[of "{x}"] assms by auto
  22.359 @@ -818,9 +818,9 @@
  22.360  
  22.361  lemma (in prob_space) measure'_translate:
  22.362    assumes X: "random_variable S X" and A: "A \<in> sets S"
  22.363 -  shows "finite_measure.\<mu>' (S\<lparr> measure := extreal\<circ>distribution X \<rparr>) A = distribution X A"
  22.364 +  shows "finite_measure.\<mu>' (S\<lparr> measure := ereal\<circ>distribution X \<rparr>) A = distribution X A"
  22.365  proof -
  22.366 -  interpret S: prob_space "S\<lparr> measure := extreal\<circ>distribution X \<rparr>"
  22.367 +  interpret S: prob_space "S\<lparr> measure := ereal\<circ>distribution X \<rparr>"
  22.368      using distribution_prob_space[OF X] .
  22.369    from A show "S.\<mu>' A = distribution X A"
  22.370      unfolding S.\<mu>'_def by (simp add: distribution_def_raw \<mu>'_def)
  22.371 @@ -831,9 +831,9 @@
  22.372    assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
  22.373    shows "\<H>(X) = log b (real (card (X ` space M)))"
  22.374  proof -
  22.375 -  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := extreal\<circ>distribution X\<rparr>"
  22.376 +  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := ereal\<circ>distribution X\<rparr>"
  22.377    note frv = simple_function_imp_finite_random_variable[OF X]
  22.378 -  from distribution_finite_prob_space[OF this, of "\<lparr> measure = extreal\<circ>distribution X \<rparr>"]
  22.379 +  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
  22.380    interpret X: finite_prob_space ?X by simp
  22.381    note rv = finite_random_variableD[OF frv]
  22.382    have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
  22.383 @@ -917,9 +917,9 @@
  22.384  abbreviation (in information_space)
  22.385    conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
  22.386    "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
  22.387 -    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
  22.388 -    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr>
  22.389 -    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = extreal\<circ>distribution Z \<rparr>
  22.390 +    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
  22.391 +    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr>
  22.392 +    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = ereal\<circ>distribution Z \<rparr>
  22.393      X Y Z"
  22.394  
  22.395  lemma (in information_space) conditional_mutual_information_generic_eq:
  22.396 @@ -1118,8 +1118,8 @@
  22.397  abbreviation (in information_space)
  22.398    conditional_entropy_Pow ("\<H>'(_ | _')") where
  22.399    "\<H>(X | Y) \<equiv> conditional_entropy b
  22.400 -    \<lparr> space = X`space M, sets = Pow (X`space M), measure = extreal\<circ>distribution X \<rparr>
  22.401 -    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = extreal\<circ>distribution Y \<rparr> X Y"
  22.402 +    \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr>
  22.403 +    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
  22.404  
  22.405  lemma (in information_space) conditional_entropy_positive:
  22.406    "simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
    23.1 --- a/src/HOL/Probability/Lebesgue_Integration.thy	Wed Jul 20 10:11:08 2011 +0200
    23.2 +++ b/src/HOL/Probability/Lebesgue_Integration.thy	Wed Jul 20 10:48:00 2011 +0200
    23.3 @@ -9,10 +9,10 @@
    23.4  imports Measure Borel_Space
    23.5  begin
    23.6  
    23.7 -lemma real_extreal_1[simp]: "real (1::extreal) = 1"
    23.8 -  unfolding one_extreal_def by simp
    23.9 +lemma real_ereal_1[simp]: "real (1::ereal) = 1"
   23.10 +  unfolding one_ereal_def by simp
   23.11  
   23.12 -lemma extreal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::extreal)"
   23.13 +lemma ereal_indicator_pos[simp,intro]: "0 \<le> (indicator A x ::ereal)"
   23.14    unfolding indicator_def by auto
   23.15  
   23.16  lemma tendsto_real_max:
   23.17 @@ -150,7 +150,7 @@
   23.18  qed
   23.19  
   23.20  lemma (in sigma_algebra) simple_function_indicator_representation:
   23.21 -  fixes f ::"'a \<Rightarrow> extreal"
   23.22 +  fixes f ::"'a \<Rightarrow> ereal"
   23.23    assumes f: "simple_function M f" and x: "x \<in> space M"
   23.24    shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
   23.25    (is "?l = ?r")
   23.26 @@ -165,7 +165,7 @@
   23.27  qed
   23.28  
   23.29  lemma (in measure_space) simple_function_notspace:
   23.30 -  "simple_function M (\<lambda>x. h x * indicator (- space M) x::extreal)" (is "simple_function M ?h")
   23.31 +  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h")
   23.32  proof -
   23.33    have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
   23.34    hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
   23.35 @@ -212,7 +212,7 @@
   23.36    by (auto intro: borel_measurable_vimage)
   23.37  
   23.38  lemma (in sigma_algebra) simple_function_eq_borel_measurable:
   23.39 -  fixes f :: "'a \<Rightarrow> extreal"
   23.40 +  fixes f :: "'a \<Rightarrow> ereal"
   23.41    shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
   23.42    using simple_function_borel_measurable[of f]
   23.43      borel_measurable_simple_function[of f]
   23.44 @@ -300,7 +300,7 @@
   23.45  
   23.46  lemma (in sigma_algebra)
   23.47    fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
   23.48 -  shows simple_function_extreal[intro, simp]: "simple_function M (\<lambda>x. extreal (f x))"
   23.49 +  shows simple_function_ereal[intro, simp]: "simple_function M (\<lambda>x. ereal (f x))"
   23.50    by (auto intro!: simple_function_compose1[OF sf])
   23.51  
   23.52  lemma (in sigma_algebra)
   23.53 @@ -309,7 +309,7 @@
   23.54    by (auto intro!: simple_function_compose1[OF sf])
   23.55  
   23.56  lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
   23.57 -  fixes u :: "'a \<Rightarrow> extreal"
   23.58 +  fixes u :: "'a \<Rightarrow> ereal"
   23.59    assumes u: "u \<in> borel_measurable M"
   23.60    shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
   23.61               (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)"
   23.62 @@ -331,7 +331,7 @@
   23.63      "\<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)))"
   23.64      unfolding f_def by auto
   23.65  
   23.66 -  let "?g j x" = "real (f x j) / 2^j :: extreal"
   23.67 +  let "?g j x" = "real (f x j) / 2^j :: ereal"
   23.68    show ?thesis
   23.69    proof (intro exI[of _ ?g] conjI allI ballI)
   23.70      fix i
   23.71 @@ -345,22 +345,22 @@
   23.72          by (rule finite_subset) auto
   23.73      qed
   23.74      then show "simple_function M (?g i)"
   23.75 -      by (auto intro: simple_function_extreal simple_function_div)
   23.76 +      by (auto intro: simple_function_ereal simple_function_div)
   23.77    next
   23.78      show "incseq ?g"
   23.79 -    proof (intro incseq_extreal incseq_SucI le_funI)
   23.80 +    proof (intro incseq_ereal incseq_SucI le_funI)
   23.81        fix x and i :: nat
   23.82        have "f x i * 2 \<le> f x (Suc i)" unfolding f_def
   23.83        proof ((split split_if)+, intro conjI impI)
   23.84 -        assume "extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
   23.85 +        assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
   23.86          then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)"
   23.87            by (cases "u x") (auto intro!: le_natfloor)
   23.88        next
   23.89 -        assume "\<not> extreal (real i) \<le> u x" "extreal (real (Suc i)) \<le> u x"
   23.90 +        assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x"
   23.91          then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i"
   23.92            by (cases "u x") auto
   23.93        next
   23.94 -        assume "\<not> extreal (real i) \<le> u x" "\<not> extreal (real (Suc i)) \<le> u x"
   23.95 +        assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x"
   23.96          have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2"
   23.97            by simp
   23.98          also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)"
   23.99 @@ -380,7 +380,7 @@
  23.100      qed
  23.101    next
  23.102      fix x show "(SUP i. ?g i x) = max 0 (u x)"
  23.103 -    proof (rule extreal_SUPI)
  23.104 +    proof (rule ereal_SUPI)
  23.105        fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def
  23.106          by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg
  23.107                                       mult_nonpos_nonneg mult_nonneg_nonneg)
  23.108 @@ -393,7 +393,7 @@
  23.109          case (real r)
  23.110          with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq)
  23.111          from real_arch_lt[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le)
  23.112 -        then have "\<exists>p. max 0 (u x) = extreal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
  23.113 +        then have "\<exists>p. max 0 (u x) = ereal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def)
  23.114          then guess p .. note ux = this
  23.115          obtain m :: nat where m: "p < real m" using real_arch_lt ..
  23.116          have "p \<le> r"
  23.117 @@ -417,7 +417,7 @@
  23.118  qed
  23.119  
  23.120  lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
  23.121 -  fixes u :: "'a \<Rightarrow> extreal"
  23.122 +  fixes u :: "'a \<Rightarrow> ereal"
  23.123    assumes u: "u \<in> borel_measurable M"
  23.124    obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
  23.125      "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
  23.126 @@ -454,7 +454,7 @@
  23.127  qed
  23.128  
  23.129  lemma (in measure_space) simple_function_restricted:
  23.130 -  fixes f :: "'a \<Rightarrow> extreal" assumes "A \<in> sets M"
  23.131 +  fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M"
  23.132    shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
  23.133      (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
  23.134  proof -
  23.135 @@ -478,7 +478,7 @@
  23.136          using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
  23.137      next
  23.138        fix x
  23.139 -      assume "indicator A x \<noteq> (0::extreal)"
  23.140 +      assume "indicator A x \<noteq> (0::ereal)"
  23.141        then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
  23.142        moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
  23.143        ultimately show "f x = 0" by auto
  23.144 @@ -527,7 +527,7 @@
  23.145    "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
  23.146  
  23.147  syntax
  23.148 -  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
  23.149 +  "_simple_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
  23.150  
  23.151  translations
  23.152    "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
  23.153 @@ -591,7 +591,7 @@
  23.154    hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
  23.155      unfolding simple_integral_def using f sets
  23.156      by (subst setsum_Sigma[symmetric])
  23.157 -       (auto intro!: setsum_cong setsum_extreal_right_distrib)
  23.158 +       (auto intro!: setsum_cong setsum_ereal_right_distrib)
  23.159    also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
  23.160    proof -
  23.161      have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
  23.162 @@ -625,7 +625,7 @@
  23.163        simple_function_partition[OF f g]
  23.164        simple_function_partition[OF g f]
  23.165      by (subst (3) Int_commute)
  23.166 -       (auto simp add: extreal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
  23.167 +       (auto simp add: ereal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
  23.168  qed
  23.169  
  23.170  lemma (in measure_space) simple_integral_setsum[simp]:
  23.171 @@ -650,8 +650,8 @@
  23.172    with assms show ?thesis
  23.173      unfolding simple_function_partition[OF mult f(1)]
  23.174                simple_function_partition[OF f(1) mult]
  23.175 -    by (subst setsum_extreal_right_distrib)
  23.176 -       (auto intro!: extreal_0_le_mult setsum_cong simp: mult_assoc)
  23.177 +    by (subst setsum_ereal_right_distrib)
  23.178 +       (auto intro!: ereal_0_le_mult setsum_cong simp: mult_assoc)
  23.179  qed
  23.180  
  23.181  lemma (in measure_space) simple_integral_mono_AE:
  23.182 @@ -673,7 +673,7 @@
  23.183      proof (cases "f x \<le> g x")
  23.184        case True then show ?thesis
  23.185          using * assms(1,2)[THEN simple_functionD(2)]
  23.186 -        by (auto intro!: extreal_mult_right_mono)
  23.187 +        by (auto intro!: ereal_mult_right_mono)
  23.188      next
  23.189        case False
  23.190        obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
  23.191 @@ -767,7 +767,7 @@
  23.192    assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
  23.193    thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
  23.194  next
  23.195 -  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::extreal}" by auto
  23.196 +  assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::ereal}" by auto
  23.197    thus ?thesis
  23.198      using simple_integral_indicator[OF assms simple_function_const[of 1]]
  23.199      using sets_into_space[OF assms]
  23.200 @@ -778,7 +778,7 @@
  23.201    assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
  23.202    shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
  23.203  proof -
  23.204 -  have "AE x. indicator N x = (0 :: extreal)"
  23.205 +  have "AE x. indicator N x = (0 :: ereal)"
  23.206      using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
  23.207    then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
  23.208      using assms apply (intro simple_integral_cong_AE) by auto
  23.209 @@ -815,7 +815,7 @@
  23.210      by (auto simp: indicator_def split: split_if_asm)
  23.211    then show "f x * \<mu> (f -` {f x} \<inter> A) =
  23.212      f x * \<mu> (?f -` {f x} \<inter> space M)"
  23.213 -    unfolding extreal_mult_cancel_left by auto
  23.214 +    unfolding ereal_mult_cancel_left by auto
  23.215  qed
  23.216  
  23.217  lemma (in measure_space) simple_integral_subalgebra:
  23.218 @@ -872,7 +872,7 @@
  23.219    "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
  23.220  
  23.221  syntax
  23.222 -  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> extreal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> extreal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
  23.223 +  "_positive_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
  23.224  
  23.225  translations
  23.226    "\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
  23.227 @@ -917,8 +917,8 @@
  23.228      have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
  23.229      proof (intro SUP_PInfty)
  23.230        fix n :: nat
  23.231 -      let ?y = "extreal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
  23.232 -      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: extreal_divide_eq)
  23.233 +      let ?y = "ereal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
  23.234 +      have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: ereal_divide_eq)
  23.235        then have "?g ?y \<in> ?A" by (rule g_in_A)
  23.236        have "real n \<le> ?y * \<mu> ?G"
  23.237          using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
  23.238 @@ -1002,13 +1002,13 @@
  23.239    assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
  23.240    and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
  23.241    shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
  23.242 -proof (rule extreal_le_mult_one_interval)
  23.243 +proof (rule ereal_le_mult_one_interval)
  23.244    have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
  23.245      using f(3) by (auto intro!: le_SUPI2 positive_integral_positive)
  23.246    then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
  23.247    have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
  23.248      using u(3) by auto
  23.249 -  fix a :: extreal assume "0 < a" "a < 1"
  23.250 +  fix a :: ereal assume "0 < a" "a < 1"
  23.251    hence "a \<noteq> 0" by auto
  23.252    let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
  23.253    have B: "\<And>i. ?B i \<in> sets M"
  23.254 @@ -1043,7 +1043,7 @@
  23.255          assume "u x \<noteq> 0"
  23.256          with `a < 1` u_range[OF `x \<in> space M`]
  23.257          have "a * u x < 1 * u x"
  23.258 -          by (intro extreal_mult_strict_right_mono) (auto simp: image_iff)
  23.259 +          by (intro ereal_mult_strict_right_mono) (auto simp: image_iff)
  23.260          also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUPR_apply)
  23.261          finally obtain i where "a * u x < f i x" unfolding SUPR_def
  23.262            by (auto simp add: less_Sup_iff)
  23.263 @@ -1056,18 +1056,18 @@
  23.264  
  23.265    have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
  23.266      unfolding simple_integral_indicator[OF B `simple_function M u`]
  23.267 -  proof (subst SUPR_extreal_setsum, safe)
  23.268 +  proof (subst SUPR_ereal_setsum, safe)
  23.269      fix x n assume "x \<in> space M"
  23.270      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)"
  23.271 -      using B_mono B_u by (auto intro!: measure_mono extreal_mult_left_mono incseq_SucI simp: extreal_zero_le_0_iff)
  23.272 +      using B_mono B_u by (auto intro!: measure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
  23.273    next
  23.274      show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
  23.275        using measure_conv u_range B_u unfolding simple_integral_def
  23.276 -      by (auto intro!: setsum_cong SUPR_extreal_cmult[symmetric])
  23.277 +      by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric])
  23.278    qed
  23.279    moreover
  23.280    have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
  23.281 -    apply (subst SUPR_extreal_cmult[symmetric])
  23.282 +    apply (subst SUPR_ereal_cmult[symmetric])
  23.283    proof (safe intro!: SUP_mono bexI)
  23.284      fix i
  23.285      have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
  23.286 @@ -1234,7 +1234,7 @@
  23.287    have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
  23.288      using u v `0 \<le> a`
  23.289      by (auto simp: incseq_Suc_iff le_fun_def
  23.290 -             intro!: add_mono extreal_mult_left_mono simple_integral_mono)
  23.291 +             intro!: add_mono ereal_mult_left_mono simple_integral_mono)
  23.292    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)"
  23.293      using u v `0 \<le> a` by (auto simp: simple_integral_positive)
  23.294    { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
  23.295 @@ -1245,26 +1245,26 @@
  23.296    proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
  23.297      show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
  23.298        using u v  `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
  23.299 -      by (auto intro!: add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg)
  23.300 +      by (auto intro!: add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg)
  23.301      { fix x
  23.302        { 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]
  23.303            by auto }
  23.304        then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
  23.305          using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x]
  23.306 -        by (subst SUPR_extreal_cmult[symmetric, OF u(6) `0 \<le> a`])
  23.307 -           (auto intro!: SUPR_extreal_add
  23.308 -                 simp: incseq_Suc_iff le_fun_def add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) }
  23.309 +        by (subst SUPR_ereal_cmult[symmetric, OF u(6) `0 \<le> a`])
  23.310 +           (auto intro!: SUPR_ereal_add
  23.311 +                 simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg) }
  23.312      then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
  23.313        unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
  23.314 -      by (intro AE_I2) (auto split: split_max simp add: extreal_add_nonneg_nonneg)
  23.315 +      by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg)
  23.316    qed
  23.317    also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
  23.318      using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
  23.319    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)"
  23.320      unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
  23.321      unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
  23.322 -    apply (subst SUPR_extreal_cmult[symmetric, OF pos(1) `0 \<le> a`])
  23.323 -    apply (subst SUPR_extreal_add[symmetric, OF inc not_MInf]) .
  23.324 +    apply (subst SUPR_ereal_cmult[symmetric, OF pos(1) `0 \<le> a`])
  23.325 +    apply (subst SUPR_ereal_add[symmetric, OF inc not_MInf]) .
  23.326    then show ?thesis by (simp add: positive_integral_max_0)
  23.327  qed
  23.328  
  23.329 @@ -1273,7 +1273,7 @@
  23.330    shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
  23.331  proof -
  23.332    have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
  23.333 -    by (auto split: split_max simp: extreal_zero_le_0_iff)
  23.334 +    by (auto split: split_max simp: ereal_zero_le_0_iff)
  23.335    have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
  23.336      by (simp add: positive_integral_max_0)
  23.337    then show ?thesis
  23.338 @@ -1302,7 +1302,7 @@
  23.339    shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
  23.340  proof -
  23.341    have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
  23.342 -    using assms by (auto split: split_max simp: extreal_add_nonneg_nonneg)
  23.343 +    using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg)
  23.344    have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
  23.345      by (simp add: positive_integral_max_0)
  23.346    also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
  23.347 @@ -1325,7 +1325,7 @@
  23.348      case (insert i P)
  23.349      then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
  23.350        "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
  23.351 -      by (auto intro!: borel_measurable_extreal_setsum setsum_nonneg)
  23.352 +      by (auto intro!: borel_measurable_ereal_setsum setsum_nonneg)
  23.353      from positive_integral_add[OF this]
  23.354      show ?case using insert by auto
  23.355    qed simp
  23.356 @@ -1342,10 +1342,10 @@
  23.357      using positive_integral_indicator by simp
  23.358    also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
  23.359      by (auto intro!: positive_integral_mono_AE
  23.360 -      simp: indicator_def extreal_zero_le_0_iff)
  23.361 +      simp: indicator_def ereal_zero_le_0_iff)
  23.362    also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
  23.363      using assms
  23.364 -    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: extreal_zero_le_0_iff)
  23.365 +    by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: ereal_zero_le_0_iff)
  23.366    finally show ?thesis .
  23.367  qed
  23.368  
  23.369 @@ -1375,10 +1375,10 @@
  23.370    shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
  23.371  proof -
  23.372    have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
  23.373 -    using assms by (auto intro: extreal_diff_positive)
  23.374 +    using assms by (auto intro: ereal_diff_positive)
  23.375    have pos_f: "AE x. 0 \<le> f x" using mono g by auto
  23.376 -  { fix a b :: extreal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
  23.377 -      by (cases rule: extreal2_cases[of a b]) auto }
  23.378 +  { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
  23.379 +      by (cases rule: ereal2_cases[of a b]) auto }
  23.380    note * = this
  23.381    then have "AE x. f x = f x - g x + g x"
  23.382      using mono positive_integral_noteq_infinite[OF g fin] assms by auto
  23.383 @@ -1387,7 +1387,7 @@
  23.384      by (rule positive_integral_cong_AE)
  23.385    show ?thesis unfolding **
  23.386      using fin positive_integral_positive[of g]
  23.387 -    by (cases rule: extreal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
  23.388 +    by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
  23.389  qed
  23.390  
  23.391  lemma (in measure_space) positive_integral_suminf:
  23.392 @@ -1397,20 +1397,20 @@
  23.393    have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
  23.394      using assms by (auto simp: AE_all_countable)
  23.395    have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
  23.396 -    using positive_integral_positive by (rule suminf_extreal_eq_SUPR)
  23.397 +    using positive_integral_positive by (rule suminf_ereal_eq_SUPR)
  23.398    also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
  23.399      unfolding positive_integral_setsum[OF f] ..
  23.400    also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
  23.401      by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
  23.402         (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
  23.403    also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
  23.404 -    by (intro positive_integral_cong_AE) (auto simp: suminf_extreal_eq_SUPR)
  23.405 +    by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUPR)
  23.406    finally show ?thesis by simp
  23.407  qed
  23.408  
  23.409  text {* Fatou's lemma: convergence theorem on limes inferior *}
  23.410  lemma (in measure_space) positive_integral_lim_INF:
  23.411 -  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> extreal"
  23.412 +  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
  23.413    assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
  23.414    shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
  23.415  proof -
  23.416 @@ -1435,7 +1435,7 @@
  23.417    show ?thesis
  23.418    proof
  23.419      have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
  23.420 -      using u by (auto simp: extreal_zero_le_0_iff)
  23.421 +      using u by (auto simp: ereal_zero_le_0_iff)
  23.422      then show "positive M' (measure M')" unfolding M'
  23.423        using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
  23.424      show "countably_additive M' (measure M')"
  23.425 @@ -1449,7 +1449,7 @@
  23.426          by (simp add: positive_integral_suminf[OF _ pos, symmetric])
  23.427        also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
  23.428          by (intro positive_integral_cong_AE)
  23.429 -           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_extreal)
  23.430 +           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_ereal)
  23.431        also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
  23.432          unfolding suminf_indicator[OF disj] ..
  23.433        finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
  23.434 @@ -1498,7 +1498,7 @@
  23.435      { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
  23.436        then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
  23.437        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))"
  23.438 -        by (subst setsum_extreal_right_distrib) (auto simp: ac_simps)
  23.439 +        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
  23.440        also have "\<dots> = f x * G i x"
  23.441          by (simp add: indicator_def if_distrib setsum_cases)
  23.442        finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
  23.443 @@ -1521,10 +1521,10 @@
  23.444    also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
  23.445      using f G' G(2)[THEN incseq_SucD] G
  23.446      by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
  23.447 -       (auto simp: extreal_mult_left_mono le_fun_def extreal_zero_le_0_iff)
  23.448 +       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
  23.449    also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
  23.450      by (intro positive_integral_cong_AE)
  23.451 -       (auto simp add: SUPR_extreal_cmult split: split_max)
  23.452 +       (auto simp add: SUPR_ereal_cmult split: split_max)
  23.453    finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
  23.454  qed
  23.455  
  23.456 @@ -1541,16 +1541,16 @@
  23.457      with positive_integral_null_set[of ?A u] u
  23.458      show "integral\<^isup>P M u = 0" by (simp add: u_eq)
  23.459    next
  23.460 -    { fix r :: extreal and n :: nat assume gt_1: "1 \<le> real n * r"
  23.461 -      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_extreal_def)
  23.462 -      then have "0 \<le> r" by (auto simp add: extreal_zero_less_0_iff) }
  23.463 +    { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
  23.464 +      then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def)
  23.465 +      then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) }
  23.466      note gt_1 = this
  23.467      assume *: "integral\<^isup>P M u = 0"
  23.468      let "?M n" = "{x \<in> space M. 1 \<le> real (n::nat) * u x}"
  23.469      have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
  23.470      proof -
  23.471        { fix n :: nat
  23.472 -        from positive_integral_Markov_inequality[OF u pos, of ?A "extreal (real n)"]
  23.473 +        from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"]
  23.474          have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
  23.475          moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
  23.476          ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
  23.477 @@ -1566,7 +1566,7 @@
  23.478          fix n :: nat and x
  23.479          assume *: "1 \<le> real n * u x"
  23.480          also from gt_1[OF this] have "real n * u x \<le> real (Suc n) * u x"
  23.481 -          using `0 \<le> u x` by (auto intro!: extreal_mult_right_mono)
  23.482 +          using `0 \<le> u x` by (auto intro!: ereal_mult_right_mono)
  23.483          finally show "1 \<le> real (Suc n) * u x" by auto
  23.484        qed
  23.485      qed
  23.486 @@ -1579,7 +1579,7 @@
  23.487          obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
  23.488          hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto
  23.489          hence "1 \<le> real j * r" using real `0 < r` by auto
  23.490 -        thus ?thesis using `0 < r` real by (auto simp: one_extreal_def)
  23.491 +        thus ?thesis using `0 < r` real by (auto simp: one_ereal_def)
  23.492        qed (insert `0 < u x`, auto)
  23.493      qed auto
  23.494      finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
  23.495 @@ -1618,7 +1618,7 @@
  23.496  proof -
  23.497    interpret R: measure_space ?R
  23.498      by (rule restricted_measure_space) fact
  23.499 -  let "?I g x" = "g x * indicator A x :: extreal"
  23.500 +  let "?I g x" = "g x * indicator A x :: ereal"
  23.501    show ?thesis
  23.502      unfolding positive_integral_def
  23.503      unfolding simple_function_restricted[OF A]
  23.504 @@ -1675,15 +1675,15 @@
  23.505  
  23.506  definition integrable where
  23.507    "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
  23.508 -    (\<integral>\<^isup>+ x. extreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) \<noteq> \<infinity>"
  23.509 +    (\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
  23.510  
  23.511  lemma integrableD[dest]:
  23.512    assumes "integrable M f"
  23.513 -  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>"
  23.514 +  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>"
  23.515    using assms unfolding integrable_def by auto
  23.516  
  23.517  definition lebesgue_integral_def:
  23.518 -  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. extreal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. extreal (- f x) \<partial>M))"
  23.519 +  "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. ereal (- f x) \<partial>M))"
  23.520  
  23.521  syntax
  23.522    "_lebesgue_integral" :: "'a \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
  23.523 @@ -1694,13 +1694,13 @@
  23.524  lemma (in measure_space) integrableE:
  23.525    assumes "integrable M f"
  23.526    obtains r q where
  23.527 -    "(\<integral>\<^isup>+x. extreal (f x)\<partial>M) = extreal r"
  23.528 -    "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M) = extreal q"
  23.529 +    "(\<integral>\<^isup>+x. ereal (f x)\<partial>M) = ereal r"
  23.530 +    "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M) = ereal q"
  23.531      "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
  23.532    using assms unfolding integrable_def lebesgue_integral_def
  23.533 -  using positive_integral_positive[of "\<lambda>x. extreal (f x)"]
  23.534 -  using positive_integral_positive[of "\<lambda>x. extreal (-f x)"]
  23.535 -  by (cases rule: extreal2_cases[of "(\<integral>\<^isup>+x. extreal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. extreal (f x)\<partial>M)"]) auto
  23.536 +  using positive_integral_positive[of "\<lambda>x. ereal (f x)"]
  23.537 +  using positive_integral_positive[of "\<lambda>x. ereal (-f x)"]
  23.538 +  by (cases rule: ereal2_cases[of "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. ereal (f x)\<partial>M)"]) auto
  23.539  
  23.540  lemma (in measure_space) integral_cong:
  23.541    assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
  23.542 @@ -1722,8 +1722,8 @@
  23.543    assumes cong: "AE x. f x = g x"
  23.544    shows "integral\<^isup>L M f = integral\<^isup>L M g"
  23.545  proof -
  23.546 -  have *: "AE x. extreal (f x) = extreal (g x)"
  23.547 -    "AE x. extreal (- f x) = extreal (- g x)" using cong by auto
  23.548 +  have *: "AE x. ereal (f x) = ereal (g x)"
  23.549 +    "AE x. ereal (- f x) = ereal (- g x)" using cong by auto
  23.550    show ?thesis
  23.551      unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
  23.552  qed
  23.553 @@ -1733,8 +1733,8 @@
  23.554    assumes "AE x. f x = g x"
  23.555    shows "integrable M f = integrable M g"
  23.556  proof -
  23.557 -  have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (g x) \<partial>M)"
  23.558 -    "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (- g x) \<partial>M)"
  23.559 +  have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (g x) \<partial>M)"
  23.560 +    "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (- g x) \<partial>M)"
  23.561      using assms by (auto intro!: positive_integral_cong_AE)
  23.562    with assms show ?thesis
  23.563      by (auto simp: integrable_def)
  23.564 @@ -1746,11 +1746,11 @@
  23.565  
  23.566  lemma (in measure_space) integral_eq_positive_integral:
  23.567    assumes f: "\<And>x. 0 \<le> f x"
  23.568 -  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
  23.569 +  shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
  23.570  proof -
  23.571 -  { fix x have "max 0 (extreal (- f x)) = 0" using f[of x] by (simp split: split_max) }
  23.572 -  then have "0 = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)" by simp
  23.573 -  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
  23.574 +  { fix x have "max 0 (ereal (- f x)) = 0" using f[of x] by (simp split: split_max) }
  23.575 +  then have "0 = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)" by simp
  23.576 +  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
  23.577    finally show ?thesis
  23.578      unfolding lebesgue_integral_def by simp
  23.579  qed
  23.580 @@ -1762,7 +1762,7 @@
  23.581  proof -
  23.582    interpret T: measure_space M' by (rule measure_space_vimage[OF T])
  23.583    from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
  23.584 -  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
  23.585 +  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
  23.586      and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
  23.587      using f by (auto simp: comp_def)
  23.588    then show ?thesis
  23.589 @@ -1777,7 +1777,7 @@
  23.590  proof -
  23.591    interpret T: measure_space M' by (rule measure_space_vimage[OF T])
  23.592    from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
  23.593 -  have borel: "(\<lambda>x. extreal (f x)) \<in> borel_measurable M'" "(\<lambda>x. extreal (- f x)) \<in> borel_measurable M'"
  23.594 +  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
  23.595      and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
  23.596      using f by (auto simp: comp_def)
  23.597    then show ?thesis
  23.598 @@ -1795,27 +1795,27 @@
  23.599      and "integrable N g = integrable M (\<lambda>x. f x * g x)" (is ?integrable)
  23.600  proof -
  23.601    from f have ms: "measure_space (M\<lparr>measure := ?d\<rparr>)"
  23.602 -    by (intro measure_space_density[where u="\<lambda>x. extreal (f x)"]) auto
  23.603 +    by (intro measure_space_density[where u="\<lambda>x. ereal (f x)"]) auto
  23.604  
  23.605 -  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>)"
  23.606 +  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>)"
  23.607      unfolding positive_integral_max_0
  23.608      by (intro measure_space.positive_integral_cong_measure) auto
  23.609 -  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (g x)) \<partial>M)"
  23.610 +  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (g x)) \<partial>M)"
  23.611      using f g by (intro positive_integral_translated_density) auto
  23.612 -  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (f x * g x)) \<partial>M)"
  23.613 +  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (f x * g x)) \<partial>M)"
  23.614      using f by (intro positive_integral_cong_AE)
  23.615 -               (auto simp: extreal_max_0 zero_le_mult_iff split: split_max)
  23.616 +               (auto simp: ereal_max_0 zero_le_mult_iff split: split_max)
  23.617    finally have pos: "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
  23.618      by (simp add: positive_integral_max_0)
  23.619    
  23.620 -  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>)"
  23.621 +  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>)"
  23.622      unfolding positive_integral_max_0
  23.623      by (intro measure_space.positive_integral_cong_measure) auto
  23.624 -  also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (- g x)) \<partial>M)"
  23.625 +  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (- g x)) \<partial>M)"
  23.626      using f g by (intro positive_integral_translated_density) auto
  23.627 -  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (- f x * g x)) \<partial>M)"
  23.628 +  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (- f x * g x)) \<partial>M)"
  23.629      using f by (intro positive_integral_cong_AE)
  23.630 -               (auto simp: extreal_max_0 mult_le_0_iff split: split_max)
  23.631 +               (auto simp: ereal_max_0 mult_le_0_iff split: split_max)
  23.632    finally have neg: "(\<integral>\<^isup>+ x. - g x \<partial>N) = (\<integral>\<^isup>+ x. - (f x * g x) \<partial>M)"
  23.633      by (simp add: positive_integral_max_0)
  23.634  
  23.635 @@ -1846,10 +1846,10 @@
  23.636    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"
  23.637    shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
  23.638  proof -
  23.639 -  let "?f x" = "max 0 (extreal (f x))"
  23.640 -  let "?mf x" = "max 0 (extreal (- f x))"
  23.641 -  let "?u x" = "max 0 (extreal (u x))"
  23.642 -  let "?v x" = "max 0 (extreal (v x))"
  23.643 +  let "?f x" = "max 0 (ereal (f x))"
  23.644 +  let "?mf x" = "max 0 (ereal (- f x))"
  23.645 +  let "?u x" = "max 0 (ereal (u x))"
  23.646 +  let "?v x" = "max 0 (ereal (v x))"
  23.647  
  23.648    from borel_measurable_diff[of u v] integrable
  23.649    have f_borel: "?f \<in> borel_measurable M" and
  23.650 @@ -1859,9 +1859,9 @@
  23.651      "f \<in> borel_measurable M"
  23.652      by (auto simp: f_def[symmetric] integrable_def)
  23.653  
  23.654 -  have "(\<integral>\<^isup>+ x. extreal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
  23.655 +  have "(\<integral>\<^isup>+ x. ereal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
  23.656      using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
  23.657 -  moreover have "(\<integral>\<^isup>+ x. extreal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
  23.658 +  moreover have "(\<integral>\<^isup>+ x. ereal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
  23.659      using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
  23.660    ultimately show f: "integrable M f"
  23.661      using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
  23.662 @@ -1886,22 +1886,22 @@
  23.663    shows "integrable M (\<lambda>t. a * f t + g t)"
  23.664    and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
  23.665  proof -
  23.666 -  let "?f x" = "max 0 (extreal (f x))"
  23.667 -  let "?g x" = "max 0 (extreal (g x))"
  23.668 -  let "?mf x" = "max 0 (extreal (- f x))"
  23.669 -  let "?mg x" = "max 0 (extreal (- g x))"
  23.670 +  let "?f x" = "max 0 (ereal (f x))"
  23.671 +  let "?g x" = "max 0 (ereal (g x))"
  23.672 +  let "?mf x" = "max 0 (ereal (- f x))"
  23.673 +  let "?mg x" = "max 0 (ereal (- g x))"
  23.674    let "?p t" = "max 0 (a * f t) + max 0 (g t)"
  23.675    let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
  23.676  
  23.677    from assms have linear:
  23.678 -    "(\<integral>\<^isup>+ x. extreal a * ?f x + ?g x \<partial>M) = extreal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
  23.679 -    "(\<integral>\<^isup>+ x. extreal a * ?mf x + ?mg x \<partial>M) = extreal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
  23.680 +    "(\<integral>\<^isup>+ x. ereal a * ?f x + ?g x \<partial>M) = ereal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
  23.681 +    "(\<integral>\<^isup>+ x. ereal a * ?mf x + ?mg x \<partial>M) = ereal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
  23.682      by (auto intro!: positive_integral_linear simp: integrable_def)
  23.683  
  23.684 -  have *: "(\<integral>\<^isup>+x. extreal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- ?n x) \<partial>M) = 0"
  23.685 +  have *: "(\<integral>\<^isup>+x. ereal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- ?n x) \<partial>M) = 0"
  23.686      using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
  23.687 -  have **: "\<And>x. extreal a * ?f x + ?g x = max 0 (extreal (?p x))"
  23.688 -           "\<And>x. extreal a * ?mf x + ?mg x = max 0 (extreal (?n x))"
  23.689 +  have **: "\<And>x. ereal a * ?f x + ?g x = max 0 (ereal (?p x))"
  23.690 +           "\<And>x. ereal a * ?mf x + ?mg x = max 0 (ereal (?n x))"
  23.691      using `0 \<le> a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff)
  23.692  
  23.693    have "integrable M ?p" "integrable M ?n"
  23.694 @@ -1958,12 +1958,12 @@
  23.695    and mono: "AE t. f t \<le> g t"
  23.696    shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
  23.697  proof -
  23.698 -  have "AE x. extreal (f x) \<le> extreal (g x)"
  23.699 +  have "AE x. ereal (f x) \<le> ereal (g x)"
  23.700      using mono by auto
  23.701 -  moreover have "AE x. extreal (- g x) \<le> extreal (- f x)"
  23.702 +  moreover have "AE x. ereal (- g x) \<le> ereal (- f x)"
  23.703      using mono by auto
  23.704    ultimately show ?thesis using fg
  23.705 -    by (auto intro!: add_mono positive_integral_mono_AE real_of_extreal_positive_mono
  23.706 +    by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
  23.707               simp: positive_integral_positive lebesgue_integral_def diff_minus)
  23.708  qed
  23.709  
  23.710 @@ -1986,9 +1986,9 @@
  23.711    and "integrable M (indicator A)" (is ?able)
  23.712  proof -
  23.713    from `A \<in> sets M` have *:
  23.714 -    "\<And>x. extreal (indicator A x) = indicator A x"
  23.715 -    "(\<integral>\<^isup>+x. extreal (- indicator A x) \<partial>M) = 0"
  23.716 -    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_extreal_def)
  23.717 +    "\<And>x. ereal (indicator A x) = indicator A x"
  23.718 +    "(\<integral>\<^isup>+x. ereal (- indicator A x) \<partial>M) = 0"
  23.719 +    by (auto split: split_indicator simp: positive_integral_0_iff_AE one_ereal_def)
  23.720    show ?int ?able
  23.721      using assms unfolding lebesgue_integral_def integrable_def
  23.722      by (auto simp: * positive_integral_indicator borel_measurable_indicator)
  23.723 @@ -2027,8 +2027,8 @@
  23.724    assumes "integrable M f"
  23.725    shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
  23.726  proof -
  23.727 -  from assms have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>)\<partial>M) = 0"
  23.728 -    "\<And>x. extreal \<bar>f x\<bar> = max 0 (extreal (f x)) + max 0 (extreal (- f x))"
  23.729 +  from assms have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
  23.730 +    "\<And>x. ereal \<bar>f x\<bar> = max 0 (ereal (f x)) + max 0 (ereal (- f x))"
  23.731      by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
  23.732    with assms show ?thesis
  23.733      by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
  23.734 @@ -2041,8 +2041,8 @@
  23.735      and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
  23.736  proof -
  23.737    interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
  23.738 -  have "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M)"
  23.739 -       "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M)"
  23.740 +  have "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M)"
  23.741 +       "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)"
  23.742      using borel by (auto intro!: positive_integral_subalgebra N sa)
  23.743    moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
  23.744      using assms unfolding measurable_def by auto
  23.745 @@ -2057,21 +2057,21 @@
  23.746    assumes borel: "g \<in> borel_measurable M"
  23.747    shows "integrable M g"
  23.748  proof -
  23.749 -  have "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal \<bar>g x\<bar> \<partial>M)"
  23.750 +  have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
  23.751      by (auto intro!: positive_integral_mono)
  23.752 -  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
  23.753 +  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
  23.754      using f by (auto intro!: positive_integral_mono)
  23.755    also have "\<dots> < \<infinity>"
  23.756      using `integrable M f` unfolding integrable_def by auto
  23.757 -  finally have pos: "(\<integral>\<^isup>+ x. extreal (g x) \<partial>M) < \<infinity>" .
  23.758 +  finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
  23.759  
  23.760 -  have "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. extreal (\<bar>g x\<bar>) \<partial>M)"
  23.761 +  have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
  23.762      by (auto intro!: positive_integral_mono)
  23.763 -  also have "\<dots> \<le> (\<integral>\<^isup>+ x. extreal (f x) \<partial>M)"
  23.764 +  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
  23.765      using f by (auto intro!: positive_integral_mono)
  23.766    also have "\<dots> < \<infinity>"
  23.767      using `integrable M f` unfolding integrable_def by auto
  23.768 -  finally have neg: "(\<integral>\<^isup>+ x. extreal (- g x) \<partial>M) < \<infinity>" .
  23.769 +  finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
  23.770  
  23.771    from neg pos borel show ?thesis
  23.772      unfolding integrable_def by auto
  23.773 @@ -2143,31 +2143,31 @@
  23.774        by (simp add: mono_def incseq_def) }
  23.775    note pos_u = this
  23.776  
  23.777 -  have SUP_F: "\<And>x. (SUP n. extreal (f n x)) = extreal (u x)"
  23.778 +  have SUP_F: "\<And>x. (SUP n. ereal (f n x)) = ereal (u x)"
  23.779      unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
  23.780  
  23.781 -  have borel_f: "\<And>i. (\<lambda>x. extreal (f i x)) \<in> borel_measurable M"
  23.782 +  have borel_f: "\<And>i. (\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
  23.783      using i unfolding integrable_def by auto
  23.784 -  hence "(\<lambda>x. SUP i. extreal (f i x)) \<in> borel_measurable M"
  23.785 +  hence "(\<lambda>x. SUP i. ereal (f i x)) \<in> borel_measurable M"
  23.786      by auto
  23.787    hence borel_u: "u \<in> borel_measurable M"
  23.788 -    by (auto simp: borel_measurable_extreal_iff SUP_F)
  23.789 +    by (auto simp: borel_measurable_ereal_iff SUP_F)
  23.790  
  23.791 -  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. extreal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. extreal (- u x) \<partial>M) = 0"
  23.792 +  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. ereal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- u x) \<partial>M) = 0"
  23.793      using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
  23.794  
  23.795 -  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. extreal (f n x) \<partial>M) = extreal (integral\<^isup>L M (f n))"
  23.796 -    using i positive_integral_positive by (auto simp: extreal_real lebesgue_integral_def integrable_def)
  23.797 +  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M) = ereal (integral\<^isup>L M (f n))"
  23.798 +    using i positive_integral_positive by (auto simp: ereal_real lebesgue_integral_def integrable_def)
  23.799  
  23.800    have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
  23.801      using pos i by (auto simp: integral_positive)
  23.802    hence "0 \<le> x"
  23.803      using LIMSEQ_le_const[OF ilim, of 0] by auto
  23.804  
  23.805 -  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))"
  23.806 +  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))"
  23.807      by (auto intro!: positive_integral_monotone_convergence_SUP
  23.808        simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
  23.809 -  also have "\<dots> = extreal x" unfolding integral_eq
  23.810 +  also have "\<dots> = ereal x" unfolding integral_eq
  23.811    proof (rule SUP_eq_LIMSEQ[THEN iffD2])
  23.812      show "mono (\<lambda>n. integral\<^isup>L M (f n))"
  23.813        using mono i by (auto simp: mono_def intro!: integral_mono)
  23.814 @@ -2205,15 +2205,15 @@
  23.815    assumes "integrable M f"
  23.816    shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
  23.817  proof -
  23.818 -  have *: "(\<integral>\<^isup>+x. extreal (- \<bar>f x\<bar>) \<partial>M) = 0"
  23.819 +  have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
  23.820      using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
  23.821    have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
  23.822 -  hence "(\<lambda>x. extreal (\<bar>f x\<bar>)) \<in> borel_measurable M"
  23.823 -    "(\<integral>\<^isup>+ x. extreal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
  23.824 +  hence "(\<lambda>x. ereal (\<bar>f x\<bar>)) \<in> borel_measurable M"
  23.825 +    "(\<integral>\<^isup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
  23.826    from positive_integral_0_iff[OF this(1)] this(2)
  23.827    show ?thesis unfolding lebesgue_integral_def *
  23.828 -    using positive_integral_positive[of "\<lambda>x. extreal \<bar>f x\<bar>"]
  23.829 -    by (auto simp add: real_of_extreal_eq_0)
  23.830 +    using positive_integral_positive[of "\<lambda>x. ereal \<bar>f x\<bar>"]
  23.831 +    by (auto simp add: real_of_ereal_eq_0)
  23.832  qed
  23.833  
  23.834  lemma (in measure_space) positive_integral_PInf:
  23.835 @@ -2255,17 +2255,17 @@
  23.836  lemma (in measure_space) integral_real:
  23.837    "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))"
  23.838    using assms unfolding lebesgue_integral_def
  23.839 -  by (subst (1 2) positive_integral_cong_AE) (auto simp add: extreal_real)
  23.840 +  by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real)
  23.841  
  23.842  lemma (in finite_measure) lebesgue_integral_const[simp]:
  23.843    shows "integrable M (\<lambda>x. a)"
  23.844    and  "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
  23.845  proof -
  23.846    { fix a :: real assume "0 \<le> a"
  23.847 -    then have "(\<integral>\<^isup>+ x. extreal a \<partial>M) = extreal a * \<mu> (space M)"
  23.848 +    then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * \<mu> (space M)"
  23.849        by (subst positive_integral_const) auto
  23.850      moreover
  23.851 -    from `0 \<le> a` have "(\<integral>\<^isup>+ x. extreal (-a) \<partial>M) = 0"
  23.852 +    from `0 \<le> a` have "(\<integral>\<^isup>+ x. ereal (-a) \<partial>M) = 0"
  23.853        by (subst positive_integral_0_iff_AE) auto
  23.854      ultimately have "integrable M (\<lambda>x. a)" by (auto simp: integrable_def) }
  23.855    note * = this
  23.856 @@ -2282,7 +2282,7 @@
  23.857  qed
  23.858  
  23.859  lemma indicator_less[simp]:
  23.860 -  "indicator A x \<le> (indicator B x::extreal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
  23.861 +  "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
  23.862    by (simp add: indicator_def not_le)
  23.863  
  23.864  lemma (in finite_measure) integral_less_AE:
  23.865 @@ -2365,21 +2365,21 @@
  23.866      finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
  23.867    note diff_less_2w = this
  23.868  
  23.869 -  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. extreal (?diff n x) \<partial>M) =
  23.870 -    (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.871 +  have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. ereal (?diff n x) \<partial>M) =
  23.872 +    (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.873      using diff w diff_less_2w w_pos
  23.874      by (subst positive_integral_diff[symmetric])
  23.875         (auto simp: integrable_def intro!: positive_integral_cong)
  23.876  
  23.877    have "integrable M (\<lambda>x. 2 * w x)"
  23.878      using w by (auto intro: integral_cmult)
  23.879 -  hence I2w_fin: "(\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
  23.880 -    borel_2w: "(\<lambda>x. extreal (2 * w x)) \<in> borel_measurable M"
  23.881 +  hence I2w_fin: "(\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
  23.882 +    borel_2w: "(\<lambda>x. ereal (2 * w x)) \<in> borel_measurable M"
  23.883      unfolding integrable_def by auto
  23.884  
  23.885 -  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
  23.886 +  have "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
  23.887    proof cases
  23.888 -    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
  23.889 +    assume eq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
  23.890      { fix n
  23.891        have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
  23.892          using diff_less_2w[of _ n] unfolding positive_integral_max_0
  23.893 @@ -2388,53 +2388,53 @@
  23.894          using positive_integral_positive[of ?f'] eq_0 by auto }
  23.895      then show ?thesis by (simp add: Limsup_const)
  23.896    next
  23.897 -    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (extreal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
  23.898 -    have "0 = limsup (\<lambda>n. 0 :: extreal)" by (simp add: Limsup_const)
  23.899 -    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.900 +    assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
  23.901 +    have "0 = limsup (\<lambda>n. 0 :: ereal)" by (simp add: Limsup_const)
  23.902 +    also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.903        by (intro limsup_mono positive_integral_positive)
  23.904 -    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)" .
  23.905 -    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (extreal (?diff n x))) \<partial>M)"
  23.906 +    finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
  23.907 +    have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
  23.908      proof (rule positive_integral_cong)
  23.909        fix x assume x: "x \<in> space M"
  23.910 -      show "max 0 (extreal (2 * w x)) = liminf (\<lambda>n. max 0 (extreal (?diff n x)))"
  23.911 -        unfolding extreal_max_0
  23.912 -      proof (rule lim_imp_Liminf[symmetric], unfold lim_extreal)
  23.913 +      show "max 0 (ereal (2 * w x)) = liminf (\<lambda>n. max 0 (ereal (?diff n x)))"
  23.914 +        unfolding ereal_max_0
  23.915 +      proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal)
  23.916          have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
  23.917            using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
  23.918          then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
  23.919 -          by (auto intro!: tendsto_real_max simp add: lim_extreal)
  23.920 +          by (auto intro!: tendsto_real_max simp add: lim_ereal)
  23.921        qed (rule trivial_limit_sequentially)
  23.922      qed
  23.923 -    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (extreal (?diff n x)) \<partial>M)"
  23.924 +    also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
  23.925        using u'_borel w u unfolding integrable_def
  23.926        by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
  23.927 -    also have "\<dots> = (\<integral>\<^isup>+ x. extreal (2 * w x) \<partial>M) -
  23.928 -        limsup (\<lambda>n. \<integral>\<^isup>+ x. extreal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.929 +    also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
  23.930 +        limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
  23.931        unfolding PI_diff positive_integral_max_0
  23.932 -      using positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"]
  23.933 -      by (subst liminf_extreal_cminus) auto
  23.934 +      using positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"]
  23.935 +      by (subst liminf_ereal_cminus) auto
  23.936      finally show ?thesis
  23.937 -      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. extreal (2 * w x)"] pos
  23.938 +      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"] pos
  23.939        unfolding positive_integral_max_0
  23.940 -      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)"])
  23.941 +      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)"])
  23.942           auto
  23.943    qed
  23.944  
  23.945    have "liminf ?f \<le> limsup ?f"
  23.946 -    by (intro extreal_Liminf_le_Limsup trivial_limit_sequentially)
  23.947 +    by (intro ereal_Liminf_le_Limsup trivial_limit_sequentially)
  23.948    moreover
  23.949 -  { have "0 = liminf (\<lambda>n. 0 :: extreal)" by (simp add: Liminf_const)
  23.950 +  { have "0 = liminf (\<lambda>n. 0 :: ereal)" by (simp add: Liminf_const)
  23.951      also have "\<dots> \<le> liminf ?f"
  23.952        by (intro liminf_mono positive_integral_positive)
  23.953      finally have "0 \<le> liminf ?f" . }
  23.954 -  ultimately have liminf_limsup_eq: "liminf ?f = extreal 0" "limsup ?f = extreal 0"
  23.955 +  ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
  23.956      using `limsup ?f = 0` by auto
  23.957 -  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)"
  23.958 +  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)"
  23.959      using diff positive_integral_positive
  23.960 -    by (subst integral_eq_positive_integral) (auto simp: extreal_real integrable_def)
  23.961 +    by (subst integral_eq_positive_integral) (auto simp: ereal_real integrable_def)
  23.962    then show ?lim_diff
  23.963 -    using extreal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
  23.964 -    by (simp add: lim_extreal)
  23.965 +    using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
  23.966 +    by (simp add: lim_ereal)
  23.967  
  23.968    show ?lim
  23.969    proof (rule LIMSEQ_I)
  23.970 @@ -2554,7 +2554,7 @@
  23.971      also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
  23.972        using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
  23.973      finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
  23.974 -      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_extreal_pos measurable_sets) }
  23.975 +      using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) }
  23.976    note int_abs_F = this
  23.977  
  23.978    have 1: "\<And>i. integrable M (\<lambda>x. ?F i x)"
  23.979 @@ -2618,15 +2618,15 @@
  23.980    and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
  23.981  proof -
  23.982    have *:
  23.983 -    "(\<integral>\<^isup>+ x. max 0 (extreal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (f x)) * \<mu> {x})"
  23.984 -    "(\<integral>\<^isup>+ x. max 0 (extreal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (extreal (- f x)) * \<mu> {x})"
  23.985 +    "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (f x)) * \<mu> {x})"
  23.986 +    "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (- f x)) * \<mu> {x})"
  23.987      by (simp_all add: positive_integral_finite_eq_setsum)
  23.988    then show "integrable M f" using finite_space finite_measure
  23.989      by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
  23.990               split: split_max)
  23.991    show ?I using finite_measure *
  23.992      apply (simp add: positive_integral_max_0 lebesgue_integral_def)
  23.993 -    apply (subst (1 2) setsum_real_of_extreal[symmetric])
  23.994 +    apply (subst (1 2) setsum_real_of_ereal[symmetric])
  23.995      apply (simp_all split: split_max add: setsum_subtractf[symmetric])
  23.996      apply (intro setsum_cong[OF refl])
  23.997      apply (simp split: split_max)
    24.1 --- a/src/HOL/Probability/Lebesgue_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    24.2 +++ b/src/HOL/Probability/Lebesgue_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    24.3 @@ -55,7 +55,7 @@
    24.4  definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
    24.5    "lebesgue = \<lparr> space = UNIV,
    24.6      sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n},
    24.7 -    measure = \<lambda>A. SUP n. extreal (integral (cube n) (indicator A)) \<rparr>"
    24.8 +    measure = \<lambda>A. SUP n. ereal (integral (cube n) (indicator A)) \<rparr>"
    24.9  
   24.10  lemma space_lebesgue[simp]: "space lebesgue = UNIV"
   24.11    unfolding lebesgue_def by simp
   24.12 @@ -141,15 +141,15 @@
   24.13        by (auto dest: lebesgueD)
   24.14      show "(\<Sum>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
   24.15      proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI)
   24.16 -      fix i n show "extreal (?m n i) \<le> extreal (?m (Suc n) i)"
   24.17 +      fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)"
   24.18          using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
   24.19      next
   24.20 -      fix i n show "0 \<le> extreal (?m n i)"
   24.21 +      fix i n show "0 \<le> ereal (?m n i)"
   24.22          using rA unfolding lebesgue_def
   24.23          by (auto intro!: le_SUPI2 integral_nonneg)
   24.24      next
   24.25 -      show "(SUP n. \<Sum>i. extreal (?m n i)) = (SUP n. extreal (?M n UNIV))"
   24.26 -      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_extreal[THEN iffD2] sums_def[THEN iffD2])
   24.27 +      show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))"
   24.28 +      proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2])
   24.29          fix n
   24.30          have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
   24.31          from lebesgueD[OF this]
   24.32 @@ -235,7 +235,7 @@
   24.33  
   24.34  lemma lmeasure_iff_LIMSEQ:
   24.35    assumes "A \<in> sets lebesgue" "0 \<le> m"
   24.36 -  shows "lebesgue.\<mu> A = extreal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
   24.37 +  shows "lebesgue.\<mu> A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
   24.38  proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ)
   24.39    show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
   24.40      using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
   24.41 @@ -261,7 +261,7 @@
   24.42  
   24.43  lemma lmeasure_finite_has_integral:
   24.44    fixes s :: "'a::ordered_euclidean_space set"
   24.45 -  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = extreal m" "0 \<le> m"
   24.46 +  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = ereal m" "0 \<le> m"
   24.47    shows "(indicator s has_integral m) UNIV"
   24.48  proof -
   24.49    let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
   24.50 @@ -324,7 +324,7 @@
   24.51  qed
   24.52  
   24.53  lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
   24.54 -  shows "lebesgue.\<mu> s = extreal m"
   24.55 +  shows "lebesgue.\<mu> s = ereal m"
   24.56  proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
   24.57    let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
   24.58    show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
   24.59 @@ -349,28 +349,28 @@
   24.60  qed
   24.61  
   24.62  lemma has_integral_iff_lmeasure:
   24.63 -  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m)"
   24.64 +  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m)"
   24.65  proof
   24.66    assume "(indicator A has_integral m) UNIV"
   24.67    with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
   24.68 -  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
   24.69 +  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
   24.70      by (auto intro: has_integral_nonneg)
   24.71  next
   24.72 -  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = extreal m"
   24.73 +  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
   24.74    then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
   24.75  qed
   24.76  
   24.77  lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
   24.78 -  shows "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))"
   24.79 +  shows "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))"
   24.80    using assms unfolding integrable_on_def
   24.81  proof safe
   24.82    fix y :: real assume "(indicator s has_integral y) UNIV"
   24.83    from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
   24.84 -  show "lebesgue.\<mu> s = extreal (integral UNIV (indicator s))" by simp
   24.85 +  show "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))" by simp
   24.86  qed
   24.87  
   24.88  lemma lebesgue_simple_function_indicator:
   24.89 -  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
   24.90 +  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   24.91    assumes f:"simple_function lebesgue f"
   24.92    shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
   24.93    by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
   24.94 @@ -427,14 +427,14 @@
   24.95          using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc)
   24.96        finally show ?thesis .
   24.97      qed }
   24.98 -  ultimately show "extreal (real n) \<le> extreal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
   24.99 +  ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))"
  24.100      using integral_const DIM_positive[where 'a='a]
  24.101      by (auto simp: cube_def content_closed_interval_cases setprod_constant)
  24.102  qed simp
  24.103  
  24.104  lemma
  24.105    fixes a b ::"'a::ordered_euclidean_space"
  24.106 -  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = extreal (content {a..b})"
  24.107 +  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = ereal (content {a..b})"
  24.108  proof -
  24.109    have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
  24.110      unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw)
  24.111 @@ -462,7 +462,7 @@
  24.112  lemma
  24.113    fixes a b :: real
  24.114    shows lmeasure_real_greaterThanAtMost[simp]:
  24.115 -    "lebesgue.\<mu> {a <.. b} = extreal (if a \<le> b then b - a else 0)"
  24.116 +    "lebesgue.\<mu> {a <.. b} = ereal (if a \<le> b then b - a else 0)"
  24.117  proof cases
  24.118    assume "a < b"
  24.119    then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
  24.120 @@ -474,7 +474,7 @@
  24.121  lemma
  24.122    fixes a b :: real
  24.123    shows lmeasure_real_atLeastLessThan[simp]:
  24.124 -    "lebesgue.\<mu> {a ..< b} = extreal (if a \<le> b then b - a else 0)"
  24.125 +    "lebesgue.\<mu> {a ..< b} = ereal (if a \<le> b then b - a else 0)"
  24.126  proof cases
  24.127    assume "a < b"
  24.128    then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
  24.129 @@ -486,7 +486,7 @@
  24.130  lemma
  24.131    fixes a b :: real
  24.132    shows lmeasure_real_greaterThanLessThan[simp]:
  24.133 -    "lebesgue.\<mu> {a <..< b} = extreal (if a \<le> b then b - a else 0)"
  24.134 +    "lebesgue.\<mu> {a <..< b} = ereal (if a \<le> b then b - a else 0)"
  24.135  proof cases
  24.136    assume "a < b"
  24.137    then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
  24.138 @@ -553,7 +553,7 @@
  24.139  qed simp
  24.140  
  24.141  lemma simple_function_has_integral:
  24.142 -  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
  24.143 +  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
  24.144    assumes f:"simple_function lebesgue f"
  24.145    and f':"range f \<subseteq> {0..<\<infinity>}"
  24.146    and om:"\<And>x. x \<in> range f \<Longrightarrow> lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
  24.147 @@ -563,16 +563,16 @@
  24.148    let "?M x" = "lebesgue.\<mu> (f -` {x} \<inter> UNIV)"
  24.149    let "?F x" = "indicator (f -` {x})"
  24.150    { fix x y assume "y \<in> range f"
  24.151 -    from subsetD[OF f' this] have "y * ?F y x = extreal (real y * ?F y x)"
  24.152 -      by (cases rule: extreal2_cases[of y "?F y x"])
  24.153 -         (auto simp: indicator_def one_extreal_def split: split_if_asm) }
  24.154 +    from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)"
  24.155 +      by (cases rule: ereal2_cases[of y "?F y x"])
  24.156 +         (auto simp: indicator_def one_ereal_def split: split_if_asm) }
  24.157    moreover
  24.158    { fix x assume x: "x\<in>range f"
  24.159      have "x * ?M x = real x * real (?M x)"
  24.160      proof cases
  24.161        assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
  24.162        with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis
  24.163 -        by (cases rule: extreal2_cases[of x "?M x"]) auto
  24.164 +        by (cases rule: ereal2_cases[of x "?M x"]) auto
  24.165      qed simp }
  24.166    ultimately
  24.167    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>
  24.168 @@ -580,12 +580,12 @@
  24.169      by simp
  24.170    also have \<dots>
  24.171    proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
  24.172 -               real_of_extreal_pos lebesgue.positive_measure ballI)
  24.173 +               real_of_ereal_pos lebesgue.positive_measure ballI)
  24.174      show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue" "\<And>y. f -` {y} \<inter> UNIV \<in> sets lebesgue"
  24.175        using lebesgue.simple_functionD[OF f] by auto
  24.176      fix y assume "real y \<noteq> 0" "y \<in> range f"
  24.177 -    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = extreal (real (?M y))"
  24.178 -      by (auto simp: extreal_real)
  24.179 +    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = ereal (real (?M y))"
  24.180 +      by (auto simp: ereal_real)
  24.181    qed
  24.182    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" .
  24.183  qed fact
  24.184 @@ -595,7 +595,7 @@
  24.185    using assms by auto
  24.186  
  24.187  lemma simple_function_has_integral':
  24.188 -  fixes f::"'a::ordered_euclidean_space \<Rightarrow> extreal"
  24.189 +  fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
  24.190    assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"
  24.191    and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
  24.192    shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
  24.193 @@ -629,7 +629,7 @@
  24.194  qed
  24.195  
  24.196  lemma positive_integral_has_integral:
  24.197 -  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
  24.198 +  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
  24.199    assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
  24.200    shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
  24.201  proof -
  24.202 @@ -648,23 +648,23 @@
  24.203    note u_fin = this
  24.204    then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
  24.205      by (rule simple_function_has_integral'[OF u(1,5)])
  24.206 -  have "\<forall>x. \<exists>r\<ge>0. f x = extreal r"
  24.207 +  have "\<forall>x. \<exists>r\<ge>0. f x = ereal r"
  24.208    proof
  24.209      fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)
  24.210 -    then show "\<exists>r\<ge>0. f x = extreal r" by (cases "f x") auto
  24.211 +    then show "\<exists>r\<ge>0. f x = ereal r" by (cases "f x") auto
  24.212    qed
  24.213 -  from choice[OF this] obtain f' where f': "f = (\<lambda>x. extreal (f' x))" "\<And>x. 0 \<le> f' x" by auto
  24.214 +  from choice[OF this] obtain f' where f': "f = (\<lambda>x. ereal (f' x))" "\<And>x. 0 \<le> f' x" by auto
  24.215  
  24.216 -  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
  24.217 +  have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
  24.218    proof
  24.219 -    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = extreal (r x)"
  24.220 +    fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"
  24.221      proof (intro choice allI)
  24.222        fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis
  24.223 -      then show "\<exists>r\<ge>0. u i x = extreal r" using u(5)[of i x] by (cases "u i x") auto
  24.224 +      then show "\<exists>r\<ge>0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto
  24.225      qed
  24.226    qed
  24.227    from choice[OF this] obtain u' where
  24.228 -      u': "u = (\<lambda>i x. extreal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
  24.229 +      u': "u = (\<lambda>i x. ereal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)
  24.230  
  24.231    have convergent: "f' integrable_on UNIV \<and>
  24.232      (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"
  24.233 @@ -672,7 +672,7 @@
  24.234      show int: "\<And>k. (u' k) integrable_on UNIV"
  24.235        using u_int unfolding integrable_on_def u' by auto
  24.236      show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)
  24.237 -      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_extreal_positive_mono)
  24.238 +      by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono)
  24.239      show "\<And>x. (\<lambda>k. u' k x) ----> f' x"
  24.240        using SUP_eq u(2)
  24.241        by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)
  24.242 @@ -686,16 +686,16 @@
  24.243        also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
  24.244          using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp
  24.245        also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
  24.246 -        by (auto intro!: real_of_extreal_positive_mono lebesgue.positive_integral_positive
  24.247 +        by (auto intro!: real_of_ereal_positive_mono lebesgue.positive_integral_positive
  24.248               lebesgue.positive_integral_mono le_SUPI simp: SUP_eq[symmetric])
  24.249        finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
  24.250      qed
  24.251    qed
  24.252  
  24.253 -  have "integral\<^isup>P lebesgue f = extreal (integral UNIV f')"
  24.254 +  have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
  24.255    proof (rule tendsto_unique[OF trivial_limit_sequentially])
  24.256      have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
  24.257 -      unfolding u_eq by (intro LIMSEQ_extreal_SUPR lebesgue.incseq_positive_integral u)
  24.258 +      unfolding u_eq by (intro LIMSEQ_ereal_SUPR lebesgue.incseq_positive_integral u)
  24.259      also note lebesgue.positive_integral_monotone_convergence_SUP
  24.260        [OF u(2)  lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric]
  24.261      finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
  24.262 @@ -704,12 +704,12 @@
  24.263      { fix k
  24.264        have "0 \<le> integral\<^isup>S lebesgue (u k)"
  24.265          using u by (auto intro!: lebesgue.simple_integral_positive)
  24.266 -      then have "integral\<^isup>S lebesgue (u k) = extreal (real (integral\<^isup>S lebesgue (u k)))"
  24.267 -        using u_fin by (auto simp: extreal_real) }
  24.268 +      then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"
  24.269 +        using u_fin by (auto simp: ereal_real) }
  24.270      note * = this
  24.271 -    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> extreal (integral UNIV f')"
  24.272 +    show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
  24.273        using convergent using u_int[THEN integral_unique, symmetric]
  24.274 -      by (subst *) (simp add: lim_extreal u')
  24.275 +      by (subst *) (simp add: lim_ereal u')
  24.276    qed
  24.277    then show ?thesis using convergent by (simp add: f' integrable_integral)
  24.278  qed
  24.279 @@ -719,9 +719,9 @@
  24.280    assumes f: "integrable lebesgue f"
  24.281    shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
  24.282  proof -
  24.283 -  let ?n = "\<lambda>x. real (extreal (max 0 (- f x)))" and ?p = "\<lambda>x. real (extreal (max 0 (f x)))"
  24.284 -  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: extreal_max)
  24.285 -  { fix f have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. extreal (max 0 (f x)) \<partial>lebesgue)"
  24.286 +  let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"
  24.287 +  have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)
  24.288 +  { fix f have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
  24.289        by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) }
  24.290    note eq = this
  24.291    show ?thesis
  24.292 @@ -731,8 +731,8 @@
  24.293      unfolding eq[of f] eq[of "\<lambda>x. - f x"]
  24.294      apply (safe intro!: positive_integral_has_integral)
  24.295      using integrableD[OF f]
  24.296 -    by (auto simp: zero_extreal_def[symmetric] positive_integral_max_0  split: split_max
  24.297 -             intro!: lebesgue.measurable_If lebesgue.borel_measurable_extreal)
  24.298 +    by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0  split: split_max
  24.299 +             intro!: lebesgue.measurable_If lebesgue.borel_measurable_ereal)
  24.300  qed
  24.301  
  24.302  lemma lebesgue_positive_integral_eq_borel:
  24.303 @@ -892,7 +892,7 @@
  24.304          by (simp add: interval_ne_empty eucl_le[where 'a='a])
  24.305        then have "lborel.\<mu> X = (\<Prod>x<DIM('a). lborel.\<mu> {a $$ x .. b $$ x})"
  24.306          by (auto simp: content_closed_interval eucl_le[where 'a='a]
  24.307 -                 intro!: setprod_extreal[symmetric])
  24.308 +                 intro!: setprod_ereal[symmetric])
  24.309        also have "\<dots> = measure ?P (?T X)"
  24.310          unfolding * by (subst lborel_space.measure_times) auto
  24.311        finally show ?thesis .
  24.312 @@ -917,7 +917,7 @@
  24.313    using lborel_eq_lborel_space[OF A] by simp
  24.314  
  24.315  lemma borel_fubini_positiv_integral:
  24.316 -  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> extreal"
  24.317 +  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
  24.318    assumes f: "f \<in> borel_measurable borel"
  24.319    shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
  24.320  proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _])
  24.321 @@ -982,7 +982,7 @@
  24.322  lemma lebesgue_real_affine:
  24.323    fixes X :: "real set"
  24.324    assumes "X \<in> sets borel" and "c \<noteq> 0"
  24.325 -  shows "measure lebesgue X = extreal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
  24.326 +  shows "measure lebesgue X = ereal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
  24.327      (is "_ = ?\<nu> X")
  24.328  proof -
  24.329    let ?E = "\<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::real .. b})\<rparr> :: real algebra"
  24.330 @@ -1010,7 +1010,7 @@
  24.331      show "measure_space (lborel\<lparr>measure := ?\<nu>\<rparr>)"
  24.332      proof
  24.333        show "positive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
  24.334 -        by (auto simp: positive_def intro!: extreal_0_le_mult borel.borel_measurable_sets)
  24.335 +        by (auto simp: positive_def intro!: ereal_0_le_mult borel.borel_measurable_sets)
  24.336        show "countably_additive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
  24.337        proof (simp add: countably_additive_def, safe)
  24.338          fix A :: "nat \<Rightarrow> real set" assume A: "range A \<subseteq> sets borel" "disjoint_family A"
  24.339 @@ -1026,7 +1026,7 @@
  24.340              by (rule disjoint_family_on_bisimulation) auto
  24.341          qed
  24.342          with Ai show "(\<Sum>n. ?\<nu> (A n)) = ?\<nu> (UNION UNIV A)"
  24.343 -          by (subst suminf_cmult_extreal)
  24.344 +          by (subst suminf_cmult_ereal)
  24.345               (auto simp: vimage_UN borel.borel_measurable_sets)
  24.346        qed
  24.347      qed
    25.1 --- a/src/HOL/Probability/Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    25.2 +++ b/src/HOL/Probability/Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    25.3 @@ -150,7 +150,7 @@
    25.4    finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
    25.5    then show ?thesis
    25.6      using fin `0 \<le> \<mu> s`
    25.7 -    unfolding extreal_eq_minus_iff by (auto simp: ac_simps)
    25.8 +    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
    25.9  qed
   25.10  
   25.11  lemma (in measure_space) measure_Diff:
   25.12 @@ -164,7 +164,7 @@
   25.13    also have "\<dots> = \<mu> (A - B) + \<mu> B"
   25.14      using measurable by (subst measure_additive[symmetric]) auto
   25.15    finally show "\<mu> (A - B) = \<mu> A - \<mu> B"
   25.16 -    unfolding extreal_eq_minus_iff
   25.17 +    unfolding ereal_eq_minus_iff
   25.18      using finite `0 \<le> \<mu> B` by auto
   25.19  qed
   25.20  
   25.21 @@ -207,7 +207,7 @@
   25.22    have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
   25.23      by (blast intro: range_subsetD [OF A])
   25.24    have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))"
   25.25 -    using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric])
   25.26 +    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
   25.27    also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)"
   25.28      by (rule measure_countably_additive)
   25.29         (auto simp add: disjoint_family_Suc ASuc A1 A2)
   25.30 @@ -244,7 +244,7 @@
   25.31  lemma (in measure_space) continuity_from_below_Lim:
   25.32    assumes A: "range A \<subseteq> sets M" "incseq A"
   25.33    shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)"
   25.34 -  using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A]
   25.35 +  using LIMSEQ_ereal_SUPR[OF measure_incseq, OF A]
   25.36      continuity_from_below[OF A] by simp
   25.37  
   25.38  lemma (in measure_space) measure_decseq:
   25.39 @@ -264,10 +264,10 @@
   25.40    have A0: "0 \<le> \<mu> (A 0)" using A by auto
   25.41  
   25.42    have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))"
   25.43 -    by (simp add: extreal_SUPR_uminus minus_extreal_def)
   25.44 +    by (simp add: ereal_SUPR_uminus minus_ereal_def)
   25.45    also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))"
   25.46 -    unfolding minus_extreal_def using A0 assms
   25.47 -    by (subst SUPR_extreal_add) (auto simp add: measure_decseq)
   25.48 +    unfolding minus_ereal_def using A0 assms
   25.49 +    by (subst SUPR_ereal_add) (auto simp add: measure_decseq)
   25.50    also have "\<dots> = (SUP n. \<mu> (A 0 - A n))"
   25.51      using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto
   25.52    also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
   25.53 @@ -280,7 +280,7 @@
   25.54    also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
   25.55      using A finite * by (simp, subst measure_Diff) auto
   25.56    finally show ?thesis
   25.57 -    unfolding extreal_minus_eq_minus_iff using finite A0 by auto
   25.58 +    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
   25.59  qed
   25.60  
   25.61  lemma (in measure_space) measure_insert:
   25.62 @@ -489,7 +489,7 @@
   25.63      also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
   25.64        using assms by (auto intro!: measure_subadditive)
   25.65      also have "\<dots> < \<mu> (T - S) + \<mu> S"
   25.66 -      using fin contr pos by (intro extreal_less_add) auto
   25.67 +      using fin contr pos by (intro ereal_less_add) auto
   25.68      also have "\<dots> = \<mu> (T \<union> S)"
   25.69        using assms by (subst measure_additive) auto
   25.70      also have "\<dots> \<le> \<mu> (space M)"
   25.71 @@ -1059,7 +1059,7 @@
   25.72    shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
   25.73    unfolding measure_additive[symmetric, OF measurable]
   25.74    using measurable(1,2)[THEN positive_measure]
   25.75 -  using finite by (cases rule: extreal2_cases[of "\<mu> A" "\<mu> B"]) auto
   25.76 +  using finite by (cases rule: ereal2_cases[of "\<mu> A" "\<mu> B"]) auto
   25.77  
   25.78  lemma (in measure_space) real_measure_finite_Union:
   25.79    assumes measurable:
   25.80 @@ -1067,7 +1067,7 @@
   25.81    assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>"
   25.82    shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
   25.83    using finite measurable(2)[THEN positive_measure]
   25.84 -  by (force intro!: setsum_real_of_extreal[symmetric]
   25.85 +  by (force intro!: setsum_real_of_ereal[symmetric]
   25.86              simp: measure_setsum[OF measurable, symmetric])
   25.87  
   25.88  lemma (in measure_space) real_measure_Diff:
   25.89 @@ -1088,7 +1088,7 @@
   25.90    shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
   25.91  proof -
   25.92    have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto
   25.93 -  with summable_sums[OF summable_extreal_pos, of "\<lambda>i. \<mu> (A i)"]
   25.94 +  with summable_sums[OF summable_ereal_pos, of "\<lambda>i. \<mu> (A i)"]
   25.95       measure_countably_additive[OF measurable]
   25.96    have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp
   25.97    moreover
   25.98 @@ -1096,14 +1096,14 @@
   25.99      have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)"
  25.100        using measurable by (auto intro!: measure_mono)
  25.101      moreover have "0 \<le> \<mu> (A i)" using measurable by auto
  25.102 -    ultimately have "\<mu> (A i) = extreal (real (\<mu> (A i)))"
  25.103 +    ultimately have "\<mu> (A i) = ereal (real (\<mu> (A i)))"
  25.104        using finite by (cases "\<mu> (A i)") auto }
  25.105    moreover
  25.106    have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto
  25.107 -  then have "\<mu> (\<Union>i. A i) = extreal (real (\<mu> (\<Union>i. A i)))"
  25.108 +  then have "\<mu> (\<Union>i. A i) = ereal (real (\<mu> (\<Union>i. A i)))"
  25.109      using finite by (cases "\<mu> (\<Union>i. A i)") auto
  25.110    ultimately show ?thesis
  25.111 -    unfolding sums_extreal[symmetric] by simp
  25.112 +    unfolding sums_ereal[symmetric] by simp
  25.113  qed
  25.114  
  25.115  lemma (in measure_space) real_measure_subadditive:
  25.116 @@ -1114,7 +1114,7 @@
  25.117    have "0 \<le> \<mu> (A \<union> B)" using measurable by auto
  25.118    then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
  25.119      using measure_subadditive[OF measurable] fin
  25.120 -    by (cases rule: extreal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
  25.121 +    by (cases rule: ereal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
  25.122  qed
  25.123  
  25.124  lemma (in measure_space) real_measure_setsum_singleton:
  25.125 @@ -1123,24 +1123,24 @@
  25.126    shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
  25.127    using measure_finite_singleton[OF S] fin
  25.128    using positive_measure[OF S(2)]
  25.129 -  by (force intro!: setsum_real_of_extreal[symmetric])
  25.130 +  by (force intro!: setsum_real_of_ereal[symmetric])
  25.131  
  25.132  lemma (in measure_space) real_continuity_from_below:
  25.133    assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
  25.134    shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
  25.135  proof -
  25.136    have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
  25.137 -  then have "extreal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
  25.138 -    using fin by (auto intro: extreal_real')
  25.139 +  then have "ereal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
  25.140 +    using fin by (auto intro: ereal_real')
  25.141    then show ?thesis
  25.142      using continuity_from_below_Lim[OF A]
  25.143 -    by (intro lim_real_of_extreal) simp
  25.144 +    by (intro lim_real_of_ereal) simp
  25.145  qed
  25.146  
  25.147  lemma (in measure_space) continuity_from_above_Lim:
  25.148    assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
  25.149    shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)"
  25.150 -  using LIMSEQ_extreal_INFI[OF measure_decseq, OF A]
  25.151 +  using LIMSEQ_ereal_INFI[OF measure_decseq, OF A]
  25.152    using continuity_from_above[OF A fin] by simp
  25.153  
  25.154  lemma (in measure_space) real_continuity_from_above:
  25.155 @@ -1151,11 +1151,11 @@
  25.156    moreover
  25.157    have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
  25.158      using A by (auto intro!: measure_mono)
  25.159 -  ultimately have "extreal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
  25.160 -    using fin by (auto intro: extreal_real')
  25.161 +  ultimately have "ereal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
  25.162 +    using fin by (auto intro: ereal_real')
  25.163    then show ?thesis
  25.164      using continuity_from_above_Lim[OF A fin]
  25.165 -    by (intro lim_real_of_extreal) simp
  25.166 +    by (intro lim_real_of_ereal) simp
  25.167  qed
  25.168  
  25.169  lemma (in measure_space) real_measure_countably_subadditive:
  25.170 @@ -1167,11 +1167,11 @@
  25.171      moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto
  25.172      ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto }
  25.173    moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
  25.174 -  ultimately have "extreal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. extreal (real (\<mu> (A i))))"
  25.175 -    using measure_countably_subadditive[OF A] by (auto simp: extreal_real)
  25.176 -  also have "\<dots> = extreal (\<Sum>i. real (\<mu> (A i)))"
  25.177 +  ultimately have "ereal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. ereal (real (\<mu> (A i))))"
  25.178 +    using measure_countably_subadditive[OF A] by (auto simp: ereal_real)
  25.179 +  also have "\<dots> = ereal (\<Sum>i. real (\<mu> (A i)))"
  25.180      using A
  25.181 -    by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin)
  25.182 +    by (auto intro!: sums_unique[symmetric] sums_ereal[THEN iffD2] summable_sums summable_real_of_ereal fin)
  25.183    finally show ?thesis by simp
  25.184  qed
  25.185  
  25.186 @@ -1199,14 +1199,14 @@
  25.187  definition (in finite_measure)
  25.188    "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)"
  25.189  
  25.190 -lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = extreal (\<mu>' A)"
  25.191 -  by (auto simp: \<mu>'_def extreal_real)
  25.192 +lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = ereal (\<mu>' A)"
  25.193 +  by (auto simp: \<mu>'_def ereal_real)
  25.194  
  25.195  lemma (in finite_measure) positive_measure'[simp, intro]: "0 \<le> \<mu>' A"
  25.196 -  unfolding \<mu>'_def by (auto simp: real_of_extreal_pos)
  25.197 +  unfolding \<mu>'_def by (auto simp: real_of_ereal_pos)
  25.198  
  25.199  lemma (in finite_measure) real_measure:
  25.200 -  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = extreal r"
  25.201 +  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = ereal r"
  25.202    using finite_measure[OF A] positive_measure[OF A] by (cases "\<mu> A") auto
  25.203  
  25.204  lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)"
  25.205 @@ -1215,8 +1215,8 @@
  25.206    moreover then have "\<mu> A \<le> \<mu> (space M)"
  25.207      using sets_into_space by (auto intro!: measure_mono)
  25.208    ultimately show ?thesis
  25.209 -    by (auto simp: \<mu>'_def intro!: real_of_extreal_positive_mono)
  25.210 -qed (simp add: \<mu>'_def real_of_extreal_pos)
  25.211 +    by (auto simp: \<mu>'_def intro!: real_of_ereal_positive_mono)
  25.212 +qed (simp add: \<mu>'_def real_of_ereal_pos)
  25.213  
  25.214  lemma (in finite_measure) restricted_finite_measure:
  25.215    assumes "S \<in> sets M"
  25.216 @@ -1302,8 +1302,8 @@
  25.217    note A[THEN subsetD, THEN finite_measure_eq, simp]
  25.218    note countable_UN[OF A, THEN finite_measure_eq, simp]
  25.219    from `summable (\<lambda>i. \<mu>' (A i))`
  25.220 -  have "(\<lambda>i. extreal (\<mu>' (A i))) sums extreal (\<Sum>i. \<mu>' (A i))"
  25.221 -    by (simp add: sums_extreal) (rule summable_sums)
  25.222 +  have "(\<lambda>i. ereal (\<mu>' (A i))) sums ereal (\<Sum>i. \<mu>' (A i))"
  25.223 +    by (simp add: sums_ereal) (rule summable_sums)
  25.224    from sums_unique[OF this, symmetric]
  25.225         measure_countably_subadditive[OF A]
  25.226    show ?thesis by simp
  25.227 @@ -1440,26 +1440,26 @@
  25.228  qed
  25.229  
  25.230  lemma suminf_cmult_indicator:
  25.231 -  fixes f :: "nat \<Rightarrow> extreal"
  25.232 +  fixes f :: "nat \<Rightarrow> ereal"
  25.233    assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
  25.234    shows "(\<Sum>n. f n * indicator (A n) x) = f i"
  25.235  proof -
  25.236 -  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)"
  25.237 +  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
  25.238      using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
  25.239 -  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: extreal)"
  25.240 +  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
  25.241      by (auto simp: setsum_cases)
  25.242 -  moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)"
  25.243 -  proof (rule extreal_SUPI)
  25.244 -    fix y :: extreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
  25.245 +  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
  25.246 +  proof (rule ereal_SUPI)
  25.247 +    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
  25.248      from this[of "Suc i"] show "f i \<le> y" by auto
  25.249    qed (insert assms, simp)
  25.250    ultimately show ?thesis using assms
  25.251 -    by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def)
  25.252 +    by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
  25.253  qed
  25.254  
  25.255  lemma suminf_indicator:
  25.256    assumes "disjoint_family A"
  25.257 -  shows "(\<Sum>n. indicator (A n) x :: extreal) = indicator (\<Union>i. A i) x"
  25.258 +  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
  25.259  proof cases
  25.260    assume *: "x \<in> (\<Union>i. A i)"
  25.261    then obtain i where "x \<in> A i" by auto
    26.1 --- a/src/HOL/Probability/Probability_Measure.thy	Wed Jul 20 10:11:08 2011 +0200
    26.2 +++ b/src/HOL/Probability/Probability_Measure.thy	Wed Jul 20 10:48:00 2011 +0200
    26.3 @@ -54,7 +54,7 @@
    26.4    by (auto simp: distribution_def intro!: arg_cong[where f=prob])
    26.5  
    26.6  lemma (in prob_space) prob_space: "prob (space M) = 1"
    26.7 -  using measure_space_1 unfolding \<mu>'_def by (simp add: one_extreal_def)
    26.8 +  using measure_space_1 unfolding \<mu>'_def by (simp add: one_ereal_def)
    26.9  
   26.10  lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
   26.11    using bounded_measure[of A] by (simp add: prob_space)
   26.12 @@ -242,7 +242,7 @@
   26.13  
   26.14  lemma (in prob_space) distribution_prob_space:
   26.15    assumes X: "random_variable S X"
   26.16 -  shows "prob_space (S\<lparr>measure := extreal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
   26.17 +  shows "prob_space (S\<lparr>measure := ereal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
   26.18  proof (rule prob_space_vimage)
   26.19    show "X \<in> measure_preserving M ?S"
   26.20      using X
   26.21 @@ -252,10 +252,10 @@
   26.22  qed
   26.23  
   26.24  lemma (in prob_space) AE_distribution:
   26.25 -  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := extreal \<circ> distribution X\<rparr>. Q x"
   26.26 +  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := ereal \<circ> distribution X\<rparr>. Q x"
   26.27    shows "AE x. Q (X x)"
   26.28  proof -
   26.29 -  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
   26.30 +  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>" using X by (rule distribution_prob_space)
   26.31    obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
   26.32      using assms unfolding X.almost_everywhere_def by auto
   26.33    from X[unfolded measurable_def] N show "AE x. Q (X x)"
   26.34 @@ -410,9 +410,9 @@
   26.35  
   26.36  lemma (in prob_space) distribution_eq_translated_integral:
   26.37    assumes "random_variable S X" "A \<in> sets S"
   26.38 -  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (indicator A)"
   26.39 +  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := ereal \<circ> distribution X\<rparr>) (indicator A)"
   26.40  proof -
   26.41 -  interpret S: prob_space "S\<lparr>measure := extreal \<circ> distribution X\<rparr>"
   26.42 +  interpret S: prob_space "S\<lparr>measure := ereal \<circ> distribution X\<rparr>"
   26.43      using assms(1) by (rule distribution_prob_space)
   26.44    show ?thesis
   26.45      using S.positive_integral_indicator(1)[of A] assms by simp
   26.46 @@ -548,7 +548,7 @@
   26.47  
   26.48  lemma (in prob_space) joint_distribution_prob_space:
   26.49    assumes "random_variable MX X" "random_variable MY Y"
   26.50 -  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
   26.51 +  shows "prob_space ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
   26.52    using random_variable_pairI[OF assms] by (rule distribution_prob_space)
   26.53  
   26.54  
   26.55 @@ -570,12 +570,12 @@
   26.56    assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
   26.57    shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
   26.58  proof -
   26.59 -  have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
   26.60 +  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
   26.61      using X by (intro finite_measure_eq[symmetric] in_P) auto
   26.62    also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
   26.63      using measure_times X by simp
   26.64 -  also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
   26.65 -    using X by (simp add: M.finite_measure_eq setprod_extreal)
   26.66 +  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
   26.67 +    using X by (simp add: M.finite_measure_eq setprod_ereal)
   26.68    finally show ?thesis by simp
   26.69  qed
   26.70  
   26.71 @@ -610,9 +610,9 @@
   26.72  
   26.73  lemma (in prob_space) distribution_finite_prob_space:
   26.74    assumes "finite_random_variable MX X"
   26.75 -  shows "finite_prob_space (MX\<lparr>measure := extreal \<circ> distribution X\<rparr>)"
   26.76 +  shows "finite_prob_space (MX\<lparr>measure := ereal \<circ> distribution X\<rparr>)"
   26.77  proof -
   26.78 -  interpret X: prob_space "MX\<lparr>measure := extreal \<circ> distribution X\<rparr>"
   26.79 +  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>"
   26.80      using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
   26.81    interpret MX: finite_sigma_algebra MX
   26.82      using assms by auto
   26.83 @@ -853,12 +853,12 @@
   26.84    assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)"
   26.85    shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
   26.86  proof -
   26.87 -  have "extreal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
   26.88 +  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
   26.89      using X by (intro finite_measure_eq[symmetric] in_P) auto
   26.90    also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
   26.91      using measure_finite_times X by simp
   26.92 -  also have "\<dots> = extreal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
   26.93 -    using X by (simp add: M.finite_measure_eq setprod_extreal)
   26.94 +  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
   26.95 +    using X by (simp add: M.finite_measure_eq setprod_ereal)
   26.96    finally show ?thesis by simp
   26.97  qed
   26.98  
   26.99 @@ -877,7 +877,7 @@
  26.100  lemma (in prob_space) joint_distribution_finite_prob_space:
  26.101    assumes X: "finite_random_variable MX X"
  26.102    assumes Y: "finite_random_variable MY Y"
  26.103 -  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := extreal \<circ> joint_distribution X Y\<rparr>)"
  26.104 +  shows "finite_prob_space ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
  26.105    by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
  26.106  
  26.107  lemma finite_prob_space_eq:
  26.108 @@ -1021,7 +1021,7 @@
  26.109        using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
  26.110      show "positive ?P (measure ?P)"
  26.111      proof (simp add: positive_def, safe)
  26.112 -      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_extreal_def)
  26.113 +      show "0 / \<mu> A = 0" using `\<mu> A \<noteq> 0` by (cases "\<mu> A") (auto simp: zero_ereal_def)
  26.114        fix B assume "B \<in> events"
  26.115        with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
  26.116        show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
  26.117 @@ -1035,12 +1035,12 @@
  26.118          with `A \<in> events` have "F i \<in> events" by auto }
  26.119        moreover then have "range F \<subseteq> events" by auto
  26.120        moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
  26.121 -        by (simp add: mult_commute divide_extreal_def)
  26.122 +        by (simp add: mult_commute divide_ereal_def)
  26.123        moreover have "0 \<le> inverse (\<mu> A)"
  26.124          using real_measure[OF `A \<in> events`] by auto
  26.125        ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
  26.126          using measure_countably_additive[of F] F
  26.127 -        by (auto simp: suminf_cmult_extreal)
  26.128 +        by (auto simp: suminf_cmult_ereal)
  26.129      qed
  26.130    qed
  26.131  qed
  26.132 @@ -1059,7 +1059,7 @@
  26.133  
  26.134  lemma (in finite_prob_space) finite_measure_space:
  26.135    fixes X :: "'a \<Rightarrow> 'x"
  26.136 -  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X\<rparr>"
  26.137 +  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X\<rparr>"
  26.138      (is "finite_measure_space ?S")
  26.139  proof (rule finite_measure_spaceI, simp_all)
  26.140    show "finite (X ` space M)" using finite_space by simp
  26.141 @@ -1072,13 +1072,13 @@
  26.142  qed
  26.143  
  26.144  lemma (in finite_prob_space) finite_prob_space_of_images:
  26.145 -  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = extreal \<circ> distribution X \<rparr>"
  26.146 -  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def)
  26.147 +  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X \<rparr>"
  26.148 +  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def)
  26.149  
  26.150  lemma (in finite_prob_space) finite_product_measure_space:
  26.151    fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
  26.152    assumes "finite s1" "finite s2"
  26.153 -  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = extreal \<circ> joint_distribution X Y\<rparr>"
  26.154 +  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = ereal \<circ> joint_distribution X Y\<rparr>"
  26.155      (is "finite_measure_space ?M")
  26.156  proof (rule finite_measure_spaceI, simp_all)
  26.157    show "finite (s1 \<times> s2)"
  26.158 @@ -1094,14 +1094,14 @@
  26.159  lemma (in finite_prob_space) finite_product_measure_space_of_images:
  26.160    shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
  26.161                                  sets = Pow (X ` space M \<times> Y ` space M),
  26.162 -                                measure = extreal \<circ> joint_distribution X Y \<rparr>"
  26.163 +                                measure = ereal \<circ> joint_distribution X Y \<rparr>"
  26.164    using finite_space by (auto intro!: finite_product_measure_space)
  26.165  
  26.166  lemma (in finite_prob_space) finite_product_prob_space_of_images:
  26.167    "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
  26.168 -                       measure = extreal \<circ> joint_distribution X Y \<rparr>"
  26.169 +                       measure = ereal \<circ> joint_distribution X Y \<rparr>"
  26.170    (is "finite_prob_space ?S")
  26.171 -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def)
  26.172 +proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def)
  26.173    have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
  26.174    thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
  26.175      by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
  26.176 @@ -1129,7 +1129,7 @@
  26.177    by (simp add: pborel_def)
  26.178  
  26.179  interpretation pborel: prob_space pborel
  26.180 -  by default (simp add: one_extreal_def pborel_def)
  26.181 +  by default (simp add: one_ereal_def pborel_def)
  26.182  
  26.183  lemma pborel_prob: "pborel.prob A = (if A \<in> sets borel \<and> A \<subseteq> {0 ..< 1} then real (lborel.\<mu> A) else 0)"
  26.184    unfolding pborel.\<mu>'_def by (auto simp: pborel_def)
  26.185 @@ -1171,11 +1171,11 @@
  26.186  subsection "Bernoulli space"
  26.187  
  26.188  definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV,
  26.189 -  measure = extreal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
  26.190 +  measure = ereal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
  26.191  
  26.192  interpretation bernoulli: finite_prob_space "bernoulli_space p" for p
  26.193    by (rule finite_prob_spaceI)
  26.194 -     (auto simp: bernoulli_space_def UNIV_bool one_extreal_def setsum_Un_disjoint intro!: setsum_nonneg)
  26.195 +     (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg)
  26.196  
  26.197  lemma bernoulli_measure:
  26.198    "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)"
    27.1 --- a/src/HOL/Probability/Radon_Nikodym.thy	Wed Jul 20 10:11:08 2011 +0200
    27.2 +++ b/src/HOL/Probability/Radon_Nikodym.thy	Wed Jul 20 10:48:00 2011 +0200
    27.3 @@ -23,7 +23,7 @@
    27.4      fix i have Ai: "A i \<in> sets M" using range by auto
    27.5      from measure positive_measure[OF this]
    27.6      show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
    27.7 -      by (auto intro!: extreal_dense simp: extreal_0_gt_inverse)
    27.8 +      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse)
    27.9    qed
   27.10    from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
   27.11      "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
   27.12 @@ -40,15 +40,15 @@
   27.13        fix N
   27.14        have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)"
   27.15          using positive_measure[OF `A N \<in> sets M`] n[of N]
   27.16 -        by (intro extreal_mult_right_mono) auto
   27.17 +        by (intro ereal_mult_right_mono) auto
   27.18        also have "\<dots> \<le> (1 / 2) ^ Suc N"
   27.19          using measure[of N] n[of N]
   27.20 -        by (cases rule: extreal2_cases[of "n N" "\<mu> (A N)"])
   27.21 -           (simp_all add: inverse_eq_divide power_divide one_extreal_def extreal_power_divide)
   27.22 +        by (cases rule: ereal2_cases[of "n N" "\<mu> (A N)"])
   27.23 +           (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
   27.24        finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" .
   27.25        show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp
   27.26      qed
   27.27 -    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_extreal by auto
   27.28 +    finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
   27.29    next
   27.30      { fix x assume "x \<in> space M"
   27.31        then obtain i where "x \<in> A i" using space[symmetric] by auto
   27.32 @@ -65,14 +65,14 @@
   27.33      qed
   27.34    next
   27.35      show "?h \<in> borel_measurable M" using range n
   27.36 -      by (auto intro!: borel_measurable_psuminf borel_measurable_extreal_times extreal_0_le_mult intro: less_imp_le)
   27.37 +      by (auto intro!: borel_measurable_psuminf borel_measurable_ereal_times ereal_0_le_mult intro: less_imp_le)
   27.38    qed
   27.39  qed
   27.40  
   27.41  subsection "Absolutely continuous"
   27.42  
   27.43  definition (in measure_space)
   27.44 -  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: extreal))"
   27.45 +  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: ereal))"
   27.46  
   27.47  lemma (in measure_space) absolutely_continuous_AE:
   27.48    assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
   27.49 @@ -199,7 +199,7 @@
   27.50        proof (induct n)
   27.51          case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
   27.52        next
   27.53 -        case 0 with M'.empty_measure show ?case by (simp add: zero_extreal_def)
   27.54 +        case 0 with M'.empty_measure show ?case by (simp add: zero_ereal_def)
   27.55        qed } note dA_less = this
   27.56      have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
   27.57      proof (rule incseq_SucI)
   27.58 @@ -433,7 +433,7 @@
   27.59        by (intro positive_integral_suminf[symmetric]) auto
   27.60      also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)"
   27.61        using `\<And>x. 0 \<le> f x`
   27.62 -      by (intro positive_integral_cong) (simp add: suminf_cmult_extreal suminf_indicator[OF `disjoint_family A`])
   27.63 +      by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`])
   27.64      finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" .
   27.65      moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
   27.66        using M'.measure_countably_additive A by (simp add: comp_def)
   27.67 @@ -447,14 +447,14 @@
   27.68        using A by (intro f_le_\<nu>) auto
   27.69      ultimately
   27.70      show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)"
   27.71 -      by (subst suminf_extreal_minus) (simp_all add: positive_integral_positive)
   27.72 +      by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive)
   27.73    next
   27.74      fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M"
   27.75 -      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def extreal_le_minus_iff)
   27.76 +      using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff)
   27.77    next
   27.78      show "\<nu> (space M) - (\<integral>\<^isup>+ x. ?F (space M) x \<partial>M) \<noteq> \<infinity>" (is "?a - ?b \<noteq> _")
   27.79        using positive_integral_positive[of "?F (space M)"]
   27.80 -      by (cases rule: extreal2_cases[of ?a ?b]) auto
   27.81 +      by (cases rule: ereal2_cases[of ?a ?b]) auto
   27.82    qed
   27.83    then interpret M: finite_measure ?M
   27.84      where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
   27.85 @@ -490,7 +490,7 @@
   27.86      def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
   27.87      ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
   27.88        using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"]
   27.89 -      by (cases rule: extreal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
   27.90 +      by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
   27.91           (simp_all add: field_simps)
   27.92      then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
   27.93      let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
   27.94 @@ -501,7 +501,7 @@
   27.95          using `0 \<le> b` by (auto simp: positive_def)
   27.96        show "countably_additive ?Mb (measure ?Mb)"
   27.97          using `0 \<le> b` measure_countably_additive
   27.98 -        by (auto simp: countably_additive_def suminf_cmult_extreal subset_eq)
   27.99 +        by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq)
  27.100        show "measure ?Mb (space ?Mb) \<noteq> \<infinity>"
  27.101          using b by auto
  27.102      qed
  27.103 @@ -513,7 +513,7 @@
  27.104      { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
  27.105        with *[OF this] have "b * \<mu> B \<le> ?t B"
  27.106          using M'.finite_measure b finite_measure M.positive_measure[OF B(1)]
  27.107 -        by (cases rule: extreal2_cases[of "?t B" "b * \<mu> B"]) auto }
  27.108 +        by (cases rule: ereal2_cases[of "?t B" "b * \<mu> B"]) auto }
  27.109      note bM_le_t = this
  27.110      let "?f0 x" = "f x + b * indicator A0 x"
  27.111      { fix A assume A: "A \<in> sets M"
  27.112 @@ -543,8 +543,8 @@
  27.113          by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") auto
  27.114        finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
  27.115      hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
  27.116 -      by (auto intro!: borel_measurable_indicator borel_measurable_extreal_add
  27.117 -                       borel_measurable_extreal_times extreal_add_nonneg_nonneg)
  27.118 +      by (auto intro!: borel_measurable_indicator borel_measurable_ereal_add
  27.119 +                       borel_measurable_ereal_times ereal_add_nonneg_nonneg)
  27.120      have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>"
  27.121        "b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>"
  27.122        using `A0 \<in> sets M` b
  27.123 @@ -552,12 +552,12 @@
  27.124          finite_measure_of_space M.finite_measure_of_space
  27.125        by auto
  27.126      have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
  27.127 -      using M'.finite_measure_of_space pos_t unfolding extreal_less_minus_iff
  27.128 +      using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff
  27.129        by (auto cong: positive_integral_cong)
  27.130      have  "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def
  27.131        using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
  27.132        using positive_integral_positive[of "?F (space M)"]
  27.133 -      by (cases rule: extreal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
  27.134 +      by (cases rule: ereal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
  27.135           (auto simp: field_simps mult_less_cancel_left)
  27.136      also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
  27.137        using space_less_A0 b
  27.138 @@ -586,10 +586,10 @@
  27.139      then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def
  27.140        using `A0 \<in> sets M` by auto
  27.141      then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto
  27.142 -    hence "0 < b * \<mu> A0" using b by (auto simp: extreal_zero_less_0_iff)
  27.143 +    hence "0 < b * \<mu> A0" using b by (auto simp: ereal_zero_less_0_iff)
  27.144      with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
  27.145        using `f \<in> G`
  27.146 -      by (intro extreal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive)
  27.147 +      by (intro ereal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive)
  27.148      also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
  27.149        by (simp cong: positive_integral_cong)
  27.150      finally have "?y < integral\<^isup>P M ?f0" by simp
  27.151 @@ -726,7 +726,7 @@
  27.152            using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
  27.153            using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"]
  27.154            using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono]
  27.155 -          by (cases rule: extreal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
  27.156 +          by (cases rule: ereal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
  27.157        qed }
  27.158      show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
  27.159      { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
  27.160 @@ -764,7 +764,7 @@
  27.161      \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
  27.162    proof
  27.163      fix i
  27.164 -    have indicator_eq: "\<And>f x A. (f x :: extreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
  27.165 +    have indicator_eq: "\<And>f x A. (f x :: ereal) * indicator (Q i \<inter> A) x * indicator (Q i) x
  27.166        = (f x * indicator (Q i) x) * indicator A x"
  27.167        unfolding indicator_def by auto
  27.168      have fm: "finite_measure (restricted_space (Q i))"
  27.169 @@ -804,12 +804,12 @@
  27.170      have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
  27.171      have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
  27.172        "\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x"
  27.173 -      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_extreal_times)
  27.174 +      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
  27.175      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)"
  27.176        using borel by (intro positive_integral_cong) (auto simp: indicator_def)
  27.177      also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)"
  27.178        using borel Qi Q0(1) `A \<in> sets M`
  27.179 -      by (subst positive_integral_add) (auto simp del: extreal_infty_mult
  27.180 +      by (subst positive_integral_add) (auto simp del: ereal_infty_mult
  27.181            simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
  27.182      also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)"
  27.183        by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto
  27.184 @@ -866,7 +866,7 @@
  27.185    show ?thesis
  27.186    proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
  27.187      show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
  27.188 -      using borel f_borel by (auto intro: borel_measurable_extreal_times)
  27.189 +      using borel f_borel by (auto intro: borel_measurable_ereal_times)
  27.190      show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto
  27.191      fix A assume "A \<in> sets M"
  27.192      then show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
  27.193 @@ -915,7 +915,7 @@
  27.194      finally have "AE x. f x \<le> g x"
  27.195        using pos borel positive_integral_PInf_AE[OF borel(2) g_fin]
  27.196        by (subst (asm) positive_integral_0_iff_AE)
  27.197 -         (auto split: split_indicator simp: not_less extreal_minus_le_iff) }
  27.198 +         (auto split: split_indicator simp: not_less ereal_minus_le_iff) }
  27.199    from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
  27.200    show "AE x. f x = g x" by auto
  27.201  qed
  27.202 @@ -942,24 +942,24 @@
  27.203    from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
  27.204    let ?N = "{x\<in>space M. f x \<noteq> f' x}"
  27.205    have "?N \<in> sets M" using borel by auto
  27.206 -  have *: "\<And>i x A. \<And>y::extreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
  27.207 +  have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
  27.208      unfolding indicator_def by auto
  27.209    have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
  27.210      by (intro finite_density_unique[THEN iffD1] allI)
  27.211 -       (auto intro!: borel_measurable_extreal_times f Int simp: *)
  27.212 +       (auto intro!: borel_measurable_ereal_times f Int simp: *)
  27.213    moreover have "AE x. ?f Q0 x = ?f' Q0 x"
  27.214    proof (rule AE_I')
  27.215 -    { fix f :: "'a \<Rightarrow> extreal" assume borel: "f \<in> borel_measurable M"
  27.216 +    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
  27.217          and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
  27.218 -      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
  27.219 +      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
  27.220        have "(\<Union>i. ?A i) \<in> null_sets"
  27.221        proof (rule null_sets_UN)
  27.222 -        fix i have "?A i \<in> sets M"
  27.223 +        fix i ::nat have "?A i \<in> sets M"
  27.224            using borel Q0(1) by auto
  27.225 -        have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. of_nat i * indicator (?A i) x \<partial>M)"
  27.226 +        have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
  27.227            unfolding eq[OF `?A i \<in> sets M`]
  27.228            by (auto intro!: positive_integral_mono simp: indicator_def)
  27.229 -        also have "\<dots> = of_nat i * \<mu> (?A i)"
  27.230 +        also have "\<dots> = i * \<mu> (?A i)"
  27.231            using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
  27.232          also have "\<dots> < \<infinity>"
  27.233            using `?A i \<in> sets M`[THEN finite_measure] by auto
  27.234 @@ -1067,7 +1067,7 @@
  27.235    interpret \<nu>: measure_space ?N
  27.236      where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
  27.237      and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
  27.238 -  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
  27.239 +  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M"
  27.240    { fix i j have "A i \<inter> Q j \<in> sets M"
  27.241      unfolding A_def using f Q
  27.242      apply (rule_tac Int)
  27.243 @@ -1095,7 +1095,7 @@
  27.244          case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
  27.245        next
  27.246          case (real r)
  27.247 -        with less_PInf_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by (auto simp: real_eq_of_nat)
  27.248 +        with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat)
  27.249          then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"])
  27.250        next
  27.251          case MInf with x show ?thesis
  27.252 @@ -1115,9 +1115,9 @@
  27.253      next
  27.254        case (Suc n)
  27.255        then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
  27.256 -        (\<integral>\<^isup>+x. of_nat (Suc n) * indicator (Q j) x \<partial>M)"
  27.257 +        (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
  27.258          by (auto intro!: positive_integral_mono simp: indicator_def A_def)
  27.259 -      also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
  27.260 +      also have "\<dots> = Suc n * \<mu> (Q j)"
  27.261          using Q by (auto intro!: positive_integral_cmult_indicator)
  27.262        also have "\<dots> < \<infinity>"
  27.263          using Q by (auto simp: real_eq_of_nat[symmetric])
  27.264 @@ -1180,7 +1180,7 @@
  27.265    also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
  27.266      using RN_deriv(3)[OF \<nu>]
  27.267      by (auto intro!: positive_integral_cong_pos split: split_if_asm
  27.268 -             simp: max_def extreal_mult_le_0_iff)
  27.269 +             simp: max_def ereal_mult_le_0_iff)
  27.270    finally show ?thesis .
  27.271  qed
  27.272  
  27.273 @@ -1287,18 +1287,18 @@
  27.274  proof -
  27.275    interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
  27.276    have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
  27.277 -  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
  27.278 +  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
  27.279    have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
  27.280    have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp
  27.281    { fix f :: "'a \<Rightarrow> real"
  27.282      { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>"
  27.283 -      have "extreal (real (RN_deriv M \<nu> x)) * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
  27.284 +      have "ereal (real (RN_deriv M \<nu> x)) * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
  27.285          by (simp add: mult_le_0_iff)
  27.286 -      then have "RN_deriv M \<nu> x * extreal (f x) = extreal (real (RN_deriv M \<nu> x) * f x)"
  27.287 -        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: extreal_real split: split_if_asm) }
  27.288 -    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)"
  27.289 -              "(\<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)"
  27.290 -      using RN_deriv_finite[OF \<nu>] unfolding extreal_mult_minus_right uminus_extreal.simps(1)[symmetric]
  27.291 +      then have "RN_deriv M \<nu> x * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
  27.292 +        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: ereal_real split: split_if_asm) }
  27.293 +    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)"
  27.294 +              "(\<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)"
  27.295 +      using RN_deriv_finite[OF \<nu>] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
  27.296        by (auto intro!: positive_integral_cong_AE) }
  27.297    note * = this
  27.298    show ?integral ?integrable
  27.299 @@ -1311,7 +1311,7 @@
  27.300    assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>")
  27.301    assumes ac: "absolutely_continuous \<nu>"
  27.302    obtains D where "D \<in> borel_measurable M"
  27.303 -    and "AE x. RN_deriv M \<nu> x = extreal (D x)"
  27.304 +    and "AE x. RN_deriv M \<nu> x = ereal (D x)"
  27.305      and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x"
  27.306      and "\<And>x. 0 \<le> D x"
  27.307  proof
  27.308 @@ -1342,13 +1342,13 @@
  27.309    qed
  27.310    ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>"
  27.311      using RN by (intro AE_iff_measurable[THEN iffD2]) auto
  27.312 -  then show "AE x. RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
  27.313 -    using RN(3) by (auto simp: extreal_real)
  27.314 -  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
  27.315 +  then show "AE x. RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
  27.316 +    using RN(3) by (auto simp: ereal_real)
  27.317 +  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
  27.318      using ac absolutely_continuous_AE[OF ms] by auto
  27.319  
  27.320    show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)"
  27.321 -    using RN by (auto intro: real_of_extreal_pos)
  27.322 +    using RN by (auto intro: real_of_ereal_pos)
  27.323  
  27.324    have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)"
  27.325      using RN by auto
  27.326 @@ -1357,7 +1357,7 @@
  27.327    finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0"
  27.328      using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto
  27.329    with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)"
  27.330 -    by (auto simp: zero_less_real_of_extreal le_less)
  27.331 +    by (auto simp: zero_less_real_of_ereal le_less)
  27.332  qed
  27.333  
  27.334  lemma (in sigma_finite_measure) RN_deriv_singleton:
    28.1 --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Jul 20 10:11:08 2011 +0200
    28.2 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Jul 20 10:48:00 2011 +0200
    28.3 @@ -8,7 +8,7 @@
    28.4    and not_empty[simp]: "S \<noteq> {}"
    28.5  
    28.6  definition (in finite_space) "M = \<lparr> space = S, sets = Pow S,
    28.7 -  measure = \<lambda>A. extreal (card A / card S) \<rparr>"
    28.8 +  measure = \<lambda>A. ereal (card A / card S) \<rparr>"
    28.9  
   28.10  lemma (in finite_space)
   28.11    shows space_M[simp]: "space M = S"
   28.12 @@ -23,7 +23,7 @@
   28.13  qed (auto simp: M_def divide_nonneg_nonneg)
   28.14  
   28.15  sublocale finite_space \<subseteq> information_space M 2
   28.16 -  by default (simp_all add: M_def one_extreal_def)
   28.17 +  by default (simp_all add: M_def one_ereal_def)
   28.18  
   28.19  lemma (in finite_space) \<mu>'_eq[simp]: "\<mu>' A = (if A \<subseteq> S then card A / card S else 0)"
   28.20    unfolding \<mu>'_def by (auto simp: M_def)
    29.1 --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Jul 20 10:11:08 2011 +0200
    29.2 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Jul 20 10:48:00 2011 +0200
    29.3 @@ -201,13 +201,13 @@
    29.4  lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
    29.5     by (auto intro!: setsum_nonneg)
    29.6  
    29.7 -sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
    29.8 +sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
    29.9    by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset)
   29.10  
   29.11 -sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>"
   29.12 -  by default (simp add: one_extreal_def)
   29.13 +sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
   29.14 +  by default (simp add: one_ereal_def)
   29.15  
   29.16 -sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. extreal (setsum p S)\<rparr>" b
   29.17 +sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>" b
   29.18    by default simp
   29.19  
   29.20  lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> \<mu>' A = setsum p A"