src/HOL/Probability/Positive_Infinite_Real.thy
changeset 38656 d5d342611edb
child 38705 aaee86c0e237
equal deleted inserted replaced
38655:5001ed24e129 38656:d5d342611edb
       
     1 (* Author: Johannes Hoelzl, TU Muenchen *)
       
     2 
       
     3 header {* A type for positive real numbers with infinity *}
       
     4 
       
     5 theory Positive_Infinite_Real
       
     6   imports Complex_Main Nat_Bijection Multivariate_Analysis
       
     7 begin
       
     8 
       
     9 lemma less_Sup_iff:
       
    10   fixes a :: "'x\<Colon>{complete_lattice,linorder}"
       
    11   shows "a < Sup S \<longleftrightarrow> (\<exists> x \<in> S. a < x)"
       
    12   unfolding not_le[symmetric] Sup_le_iff by auto
       
    13 
       
    14 lemma Inf_less_iff:
       
    15   fixes a :: "'x\<Colon>{complete_lattice,linorder}"
       
    16   shows "Inf S < a \<longleftrightarrow> (\<exists> x \<in> S. x < a)"
       
    17   unfolding not_le[symmetric] le_Inf_iff by auto
       
    18 
       
    19 lemma SUPR_fun_expand: "(SUP y:A. f y) = (\<lambda>x. SUP y:A. f y x)"
       
    20   unfolding SUPR_def expand_fun_eq Sup_fun_def
       
    21   by (auto intro!: arg_cong[where f=Sup])
       
    22 
       
    23 lemma real_Suc_natfloor: "r < real (Suc (natfloor r))"
       
    24 proof cases
       
    25   assume "0 \<le> r"
       
    26   from floor_correct[of r]
       
    27   have "r < real (\<lfloor>r\<rfloor> + 1)" by auto
       
    28   also have "\<dots> = real (Suc (natfloor r))"
       
    29     using `0 \<le> r` by (auto simp: real_of_nat_Suc natfloor_def)
       
    30   finally show ?thesis .
       
    31 next
       
    32   assume "\<not> 0 \<le> r"
       
    33   hence "r < 0" by auto
       
    34   also have "0 < real (Suc (natfloor r))" by auto
       
    35   finally show ?thesis .
       
    36 qed
       
    37 
       
    38 lemma (in complete_lattice) Sup_mono:
       
    39   assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
       
    40   shows "Sup A \<le> Sup B"
       
    41 proof (rule Sup_least)
       
    42   fix a assume "a \<in> A"
       
    43   with assms obtain b where "b \<in> B" and "a \<le> b" by auto
       
    44   hence "b \<le> Sup B" by (auto intro: Sup_upper)
       
    45   with `a \<le> b` show "a \<le> Sup B" by auto
       
    46 qed
       
    47 
       
    48 lemma (in complete_lattice) Inf_mono:
       
    49   assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
       
    50   shows "Inf A \<le> Inf B"
       
    51 proof (rule Inf_greatest)
       
    52   fix b assume "b \<in> B"
       
    53   with assms obtain a where "a \<in> A" and "a \<le> b" by auto
       
    54   hence "Inf A \<le> a" by (auto intro: Inf_lower)
       
    55   with `a \<le> b` show "Inf A \<le> b" by auto
       
    56 qed
       
    57 
       
    58 lemma (in complete_lattice) Sup_mono_offset:
       
    59   fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
       
    60   assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" and "0 \<le> k"
       
    61   shows "(SUP n . f (k + n)) = (SUP n. f n)"
       
    62 proof (rule antisym)
       
    63   show "(SUP n. f (k + n)) \<le> (SUP n. f n)"
       
    64     by (auto intro!: Sup_mono simp: SUPR_def)
       
    65   { fix n :: 'b
       
    66     have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
       
    67     with * have "f n \<le> f (k + n)" by simp }
       
    68   thus "(SUP n. f n) \<le> (SUP n. f (k + n))"
       
    69     by (auto intro!: Sup_mono exI simp: SUPR_def)
       
    70 qed
       
    71 
       
    72 lemma (in complete_lattice) Sup_mono_offset_Suc:
       
    73   assumes *: "\<And>x. f x \<le> f (Suc x)"
       
    74   shows "(SUP n . f (Suc n)) = (SUP n. f n)"
       
    75   unfolding Suc_eq_plus1
       
    76   apply (subst add_commute)
       
    77   apply (rule Sup_mono_offset)
       
    78   by (auto intro!: order.lift_Suc_mono_le[of "op \<le>" "op <" f, OF _ *]) default
       
    79 
       
    80 lemma (in complete_lattice) Inf_start:
       
    81   assumes *: "\<And>x. f 0 \<le> f x"
       
    82   shows "(INF n. f n) = f 0"
       
    83 proof (rule antisym)
       
    84   show "(INF n. f n) \<le> f 0" by (rule INF_leI) simp
       
    85   show "f 0 \<le> (INF n. f n)" by (rule le_INFI[OF *])
       
    86 qed
       
    87 
       
    88 lemma (in complete_lattice) isotone_converge:
       
    89   fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y "
       
    90   shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
       
    91 proof -
       
    92   have "\<And>n. (SUP m. f (n + m)) = (SUP n. f n)"
       
    93     apply (rule Sup_mono_offset)
       
    94     apply (rule assms)
       
    95     by simp_all
       
    96   moreover
       
    97   { fix n have "(INF m. f (n + m)) = f n"
       
    98       using Inf_start[of "\<lambda>m. f (n + m)"] assms by simp }
       
    99   ultimately show ?thesis by simp
       
   100 qed
       
   101 
       
   102 lemma (in complete_lattice) Inf_mono_offset:
       
   103   fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
       
   104   assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x" and "0 \<le> k"
       
   105   shows "(INF n . f (k + n)) = (INF n. f n)"
       
   106 proof (rule antisym)
       
   107   show "(INF n. f n) \<le> (INF n. f (k + n))"
       
   108     by (auto intro!: Inf_mono simp: INFI_def)
       
   109   { fix n :: 'b
       
   110     have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
       
   111     with * have "f (k + n) \<le> f n" by simp }
       
   112   thus "(INF n. f (k + n)) \<le> (INF n. f n)"
       
   113     by (auto intro!: Inf_mono exI simp: INFI_def)
       
   114 qed
       
   115 
       
   116 lemma (in complete_lattice) Sup_start:
       
   117   assumes *: "\<And>x. f x \<le> f 0"
       
   118   shows "(SUP n. f n) = f 0"
       
   119 proof (rule antisym)
       
   120   show "f 0 \<le> (SUP n. f n)" by (rule le_SUPI) auto
       
   121   show "(SUP n. f n) \<le> f 0" by (rule SUP_leI[OF *])
       
   122 qed
       
   123 
       
   124 lemma (in complete_lattice) antitone_converges:
       
   125   fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x"
       
   126   shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
       
   127 proof -
       
   128   have "\<And>n. (INF m. f (n + m)) = (INF n. f n)"
       
   129     apply (rule Inf_mono_offset)
       
   130     apply (rule assms)
       
   131     by simp_all
       
   132   moreover
       
   133   { fix n have "(SUP m. f (n + m)) = f n"
       
   134       using Sup_start[of "\<lambda>m. f (n + m)"] assms by simp }
       
   135   ultimately show ?thesis by simp
       
   136 qed
       
   137 
       
   138 text {*
       
   139 
       
   140 We introduce the the positive real numbers as needed for measure theory.
       
   141 
       
   142 *}
       
   143 
       
   144 typedef pinfreal = "(Some ` {0::real..}) \<union> {None}"
       
   145   by (rule exI[of _ None]) simp
       
   146 
       
   147 subsection "Introduce @{typ pinfreal} similar to a datatype"
       
   148 
       
   149 definition "Real x = Abs_pinfreal (Some (sup 0 x))"
       
   150 definition "\<omega> = Abs_pinfreal None"
       
   151 
       
   152 definition "pinfreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
       
   153 
       
   154 definition "of_pinfreal = pinfreal_case (\<lambda>x. x) 0"
       
   155 
       
   156 defs (overloaded)
       
   157   real_of_pinfreal_def [code_unfold]: "real == of_pinfreal"
       
   158 
       
   159 lemma pinfreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pinfreal"
       
   160   unfolding pinfreal_def by simp
       
   161 
       
   162 lemma pinfreal_Some_sup[simp]: "Some (sup 0 x) \<in> pinfreal"
       
   163   by (simp add: sup_ge1)
       
   164 
       
   165 lemma pinfreal_None[simp]: "None \<in> pinfreal"
       
   166   unfolding pinfreal_def by simp
       
   167 
       
   168 lemma Real_inj[simp]:
       
   169   assumes  "0 \<le> x" and "0 \<le> y"
       
   170   shows "Real x = Real y \<longleftrightarrow> x = y"
       
   171   unfolding Real_def assms[THEN sup_absorb2]
       
   172   using assms by (simp add: Abs_pinfreal_inject)
       
   173 
       
   174 lemma Real_neq_\<omega>[simp]:
       
   175   "Real x = \<omega> \<longleftrightarrow> False"
       
   176   "\<omega> = Real x \<longleftrightarrow> False"
       
   177   by (simp_all add: Abs_pinfreal_inject \<omega>_def Real_def)
       
   178 
       
   179 lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
       
   180   unfolding Real_def by (auto simp add: Abs_pinfreal_inject intro!: sup_absorb1)
       
   181 
       
   182 lemma pinfreal_cases[case_names preal infinite, cases type: pinfreal]:
       
   183   assumes preal: "\<And>r. x = Real r \<Longrightarrow> 0 \<le> r \<Longrightarrow> P" and inf: "x = \<omega> \<Longrightarrow> P"
       
   184   shows P
       
   185 proof (cases x rule: pinfreal.Abs_pinfreal_cases)
       
   186   case (Abs_pinfreal y)
       
   187   hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
       
   188     unfolding pinfreal_def by auto
       
   189   thus P
       
   190   proof (rule disjE)
       
   191     assume "\<exists>x\<ge>0. y = Some x" then guess x ..
       
   192     thus P by (simp add: preal[of x] Real_def Abs_pinfreal(1) sup_absorb2)
       
   193   qed (simp add: \<omega>_def Abs_pinfreal(1) inf)
       
   194 qed
       
   195 
       
   196 lemma pinfreal_case_\<omega>[simp]: "pinfreal_case f i \<omega> = i"
       
   197   unfolding pinfreal_case_def by simp
       
   198 
       
   199 lemma pinfreal_case_Real[simp]: "pinfreal_case f i (Real x) = (if 0 \<le> x then f x else f 0)"
       
   200 proof (cases "0 \<le> x")
       
   201   case True thus ?thesis unfolding pinfreal_case_def by (auto intro: theI2)
       
   202 next
       
   203   case False
       
   204   moreover have "(THE r. 0 \<le> r \<and> Real 0 = Real r) = 0"
       
   205     by (auto intro!: the_equality)
       
   206   ultimately show ?thesis unfolding pinfreal_case_def by (simp add: Real_neg)
       
   207 qed
       
   208 
       
   209 lemma pinfreal_case_cancel[simp]: "pinfreal_case (\<lambda>c. i) i x = i"
       
   210   by (cases x) simp_all
       
   211 
       
   212 lemma pinfreal_case_split:
       
   213   "P (pinfreal_case f i x) = ((x = \<omega> \<longrightarrow> P i) \<and> (\<forall>r\<ge>0. x = Real r \<longrightarrow> P (f r)))"
       
   214   by (cases x) simp_all
       
   215 
       
   216 lemma pinfreal_case_split_asm:
       
   217   "P (pinfreal_case f i x) = (\<not> (x = \<omega> \<and> \<not> P i \<or> (\<exists>r. r \<ge> 0 \<and> x = Real r \<and> \<not> P (f r))))"
       
   218   by (cases x) auto
       
   219 
       
   220 lemma pinfreal_case_cong[cong]:
       
   221   assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
       
   222   shows "pinfreal_case f i x = pinfreal_case f' i' x'"
       
   223   unfolding eq using cong by (cases x') simp_all
       
   224 
       
   225 lemma real_Real[simp]: "real (Real x) = (if 0 \<le> x then x else 0)"
       
   226   unfolding real_of_pinfreal_def of_pinfreal_def by simp
       
   227 
       
   228 lemma Real_real_image:
       
   229   assumes "\<omega> \<notin> A" shows "Real ` real ` A = A"
       
   230 proof safe
       
   231   fix x assume "x \<in> A"
       
   232   hence *: "x = Real (real x)"
       
   233     using `\<omega> \<notin> A` by (cases x) auto
       
   234   show "x \<in> Real ` real ` A"
       
   235     using `x \<in> A` by (subst *) (auto intro!: imageI)
       
   236 next
       
   237   fix x assume "x \<in> A"
       
   238   thus "Real (real x) \<in> A"
       
   239     using `\<omega> \<notin> A` by (cases x) auto
       
   240 qed
       
   241 
       
   242 lemma real_pinfreal_nonneg[simp, intro]: "0 \<le> real (x :: pinfreal)"
       
   243   unfolding real_of_pinfreal_def of_pinfreal_def
       
   244   by (cases x) auto
       
   245 
       
   246 lemma real_\<omega>[simp]: "real \<omega> = 0"
       
   247   unfolding real_of_pinfreal_def of_pinfreal_def by simp
       
   248 
       
   249 lemma pinfreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
       
   250 
       
   251 subsection "@{typ pinfreal} is a monoid for addition"
       
   252 
       
   253 instantiation pinfreal :: comm_monoid_add
       
   254 begin
       
   255 
       
   256 definition "0 = Real 0"
       
   257 definition "x + y = pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
       
   258 
       
   259 lemma pinfreal_plus[simp]:
       
   260   "Real r + Real p = (if 0 \<le> r then if 0 \<le> p then Real (r + p) else Real r else Real p)"
       
   261   "x + 0 = x"
       
   262   "0 + x = x"
       
   263   "x + \<omega> = \<omega>"
       
   264   "\<omega> + x = \<omega>"
       
   265   by (simp_all add: plus_pinfreal_def Real_neg zero_pinfreal_def split: pinfreal_case_split)
       
   266 
       
   267 lemma \<omega>_neq_0[simp]:
       
   268   "\<omega> = 0 \<longleftrightarrow> False"
       
   269   "0 = \<omega> \<longleftrightarrow> False"
       
   270   by (simp_all add: zero_pinfreal_def)
       
   271 
       
   272 lemma Real_eq_0[simp]:
       
   273   "Real r = 0 \<longleftrightarrow> r \<le> 0"
       
   274   "0 = Real r \<longleftrightarrow> r \<le> 0"
       
   275   by (auto simp add: Abs_pinfreal_inject zero_pinfreal_def Real_def sup_real_def)
       
   276 
       
   277 lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pinfreal_def)
       
   278 
       
   279 instance
       
   280 proof
       
   281   fix a :: pinfreal
       
   282   show "0 + a = a" by (cases a) simp_all
       
   283 
       
   284   fix b show "a + b = b + a"
       
   285     by (cases a, cases b) simp_all
       
   286 
       
   287   fix c show "a + b + c = a + (b + c)"
       
   288     by (cases a, cases b, cases c) simp_all
       
   289 qed
       
   290 end
       
   291 
       
   292 lemma pinfreal_plus_eq_\<omega>[simp]: "(a :: pinfreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
       
   293   by (cases a, cases b) auto
       
   294 
       
   295 lemma pinfreal_add_cancel_left:
       
   296   "a + b = a + c \<longleftrightarrow> (a = \<omega> \<or> b = c)"
       
   297   by (cases a, cases b, cases c, simp_all, cases c, simp_all)
       
   298 
       
   299 lemma pinfreal_add_cancel_right:
       
   300   "b + a = c + a \<longleftrightarrow> (a = \<omega> \<or> b = c)"
       
   301   by (cases a, cases b, cases c, simp_all, cases c, simp_all)
       
   302 
       
   303 lemma Real_eq_Real:
       
   304   "Real a = Real b \<longleftrightarrow> (a = b \<or> (a \<le> 0 \<and> b \<le> 0))"
       
   305 proof (cases "a \<le> 0 \<or> b \<le> 0")
       
   306   case False with Real_inj[of a b] show ?thesis by auto
       
   307 next
       
   308   case True
       
   309   thus ?thesis
       
   310   proof
       
   311     assume "a \<le> 0"
       
   312     hence *: "Real a = 0" by simp
       
   313     show ?thesis using `a \<le> 0` unfolding * by auto
       
   314   next
       
   315     assume "b \<le> 0"
       
   316     hence *: "Real b = 0" by simp
       
   317     show ?thesis using `b \<le> 0` unfolding * by auto
       
   318   qed
       
   319 qed
       
   320 
       
   321 lemma real_pinfreal_0[simp]: "real (0 :: pinfreal) = 0"
       
   322   unfolding zero_pinfreal_def real_Real by simp
       
   323 
       
   324 lemma real_of_pinfreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
       
   325   by (cases X) auto
       
   326 
       
   327 lemma real_of_pinfreal_eq: "real X = real Y \<longleftrightarrow>
       
   328     (X = Y \<or> (X = 0 \<and> Y = \<omega>) \<or> (Y = 0 \<and> X = \<omega>))"
       
   329   by (cases X, cases Y) (auto simp add: real_of_pinfreal_eq_0)
       
   330 
       
   331 lemma real_of_pinfreal_add: "real X + real Y =
       
   332     (if X = \<omega> then real Y else if Y = \<omega> then real X else real (X + Y))"
       
   333   by (auto simp: pinfreal_noteq_omega_Ex)
       
   334 
       
   335 subsection "@{typ pinfreal} is a monoid for multiplication"
       
   336 
       
   337 instantiation pinfreal :: comm_monoid_mult
       
   338 begin
       
   339 
       
   340 definition "1 = Real 1"
       
   341 definition "x * y = (if x = 0 \<or> y = 0 then 0 else
       
   342   pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
       
   343 
       
   344 lemma pinfreal_times[simp]:
       
   345   "Real r * Real p = (if 0 \<le> r \<and> 0 \<le> p then Real (r * p) else 0)"
       
   346   "\<omega> * x = (if x = 0 then 0 else \<omega>)"
       
   347   "x * \<omega> = (if x = 0 then 0 else \<omega>)"
       
   348   "0 * x = 0"
       
   349   "x * 0 = 0"
       
   350   "1 = \<omega> \<longleftrightarrow> False"
       
   351   "\<omega> = 1 \<longleftrightarrow> False"
       
   352   by (auto simp add: times_pinfreal_def one_pinfreal_def)
       
   353 
       
   354 lemma pinfreal_one_mult[simp]:
       
   355   "Real x + 1 = (if 0 \<le> x then Real (x + 1) else 1)"
       
   356   "1 + Real x = (if 0 \<le> x then Real (1 + x) else 1)"
       
   357   unfolding one_pinfreal_def by simp_all
       
   358 
       
   359 instance
       
   360 proof
       
   361   fix a :: pinfreal show "1 * a = a"
       
   362     by (cases a) (simp_all add: one_pinfreal_def)
       
   363 
       
   364   fix b show "a * b = b * a"
       
   365     by (cases a, cases b) (simp_all add: mult_nonneg_nonneg)
       
   366 
       
   367   fix c show "a * b * c = a * (b * c)"
       
   368     apply (cases a, cases b, cases c)
       
   369     apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
       
   370     apply (cases b, cases c)
       
   371     apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
       
   372     done
       
   373 qed
       
   374 end
       
   375 
       
   376 lemma pinfreal_mult_cancel_left:
       
   377   "a * b = a * c \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
       
   378   by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
       
   379 
       
   380 lemma pinfreal_mult_cancel_right:
       
   381   "b * a = c * a \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
       
   382   by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
       
   383 
       
   384 lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pinfreal_def)
       
   385 
       
   386 lemma real_pinfreal_1[simp]: "real (1 :: pinfreal) = 1"
       
   387   unfolding one_pinfreal_def real_Real by simp
       
   388 
       
   389 lemma real_of_pinfreal_mult: "real X * real Y = real (X * Y :: pinfreal)"
       
   390   by (cases X, cases Y) (auto simp: zero_le_mult_iff)
       
   391 
       
   392 subsection "@{typ pinfreal} is a linear order"
       
   393 
       
   394 instantiation pinfreal :: linorder
       
   395 begin
       
   396 
       
   397 definition "x < y \<longleftrightarrow> pinfreal_case (\<lambda>i. pinfreal_case (\<lambda>j. i < j) True y) False x"
       
   398 definition "x \<le> y \<longleftrightarrow> pinfreal_case (\<lambda>j. pinfreal_case (\<lambda>i. i \<le> j) False x) True y"
       
   399 
       
   400 lemma pinfreal_less[simp]:
       
   401   "Real r < \<omega>"
       
   402   "Real r < Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r < p else 0 < p)"
       
   403   "\<omega> < x \<longleftrightarrow> False"
       
   404   "0 < \<omega>"
       
   405   "0 < Real r \<longleftrightarrow> 0 < r"
       
   406   "x < 0 \<longleftrightarrow> False"
       
   407   "0 < (1::pinfreal)"
       
   408   by (simp_all add: less_pinfreal_def zero_pinfreal_def one_pinfreal_def del: Real_0 Real_1)
       
   409 
       
   410 lemma pinfreal_less_eq[simp]:
       
   411   "x \<le> \<omega>"
       
   412   "Real r \<le> Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r \<le> p else r \<le> 0)"
       
   413   "0 \<le> x"
       
   414   by (simp_all add: less_eq_pinfreal_def zero_pinfreal_def del: Real_0)
       
   415 
       
   416 lemma pinfreal_\<omega>_less_eq[simp]:
       
   417   "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
       
   418   by (cases x) (simp_all add: not_le less_eq_pinfreal_def)
       
   419 
       
   420 lemma pinfreal_less_eq_zero[simp]:
       
   421   "(x::pinfreal) \<le> 0 \<longleftrightarrow> x = 0"
       
   422   by (cases x) (simp_all add: zero_pinfreal_def del: Real_0)
       
   423 
       
   424 instance
       
   425 proof
       
   426   fix x :: pinfreal
       
   427   show "x \<le> x" by (cases x) simp_all
       
   428   fix y
       
   429   show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
       
   430     by (cases x, cases y) auto
       
   431   show "x \<le> y \<or> y \<le> x "
       
   432     by (cases x, cases y) auto
       
   433   { assume "x \<le> y" "y \<le> x" thus "x = y"
       
   434       by (cases x, cases y) auto }
       
   435   { fix z assume "x \<le> y" "y \<le> z"
       
   436     thus "x \<le> z" by (cases x, cases y, cases z) auto }
       
   437 qed
       
   438 end
       
   439 
       
   440 lemma pinfreal_zero_lessI[intro]:
       
   441   "(a :: pinfreal) \<noteq> 0 \<Longrightarrow> 0 < a"
       
   442   by (cases a) auto
       
   443 
       
   444 lemma pinfreal_less_omegaI[intro, simp]:
       
   445   "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
       
   446   by (cases a) auto
       
   447 
       
   448 lemma pinfreal_plus_eq_0[simp]: "(a :: pinfreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
       
   449   by (cases a, cases b) auto
       
   450 
       
   451 lemma pinfreal_le_add1[simp, intro]: "n \<le> n + (m::pinfreal)"
       
   452   by (cases n, cases m) simp_all
       
   453 
       
   454 lemma pinfreal_le_add2: "(n::pinfreal) + m \<le> k \<Longrightarrow> m \<le> k"
       
   455   by (cases n, cases m, cases k) simp_all
       
   456 
       
   457 lemma pinfreal_le_add3: "(n::pinfreal) + m \<le> k \<Longrightarrow> n \<le> k"
       
   458   by (cases n, cases m, cases k) simp_all
       
   459 
       
   460 lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
       
   461   by (cases x) auto
       
   462 
       
   463 subsection {* @{text "x - y"} on @{typ pinfreal} *}
       
   464 
       
   465 instantiation pinfreal :: minus
       
   466 begin
       
   467 definition "x - y = (if y < x then THE d. x = y + d else 0 :: pinfreal)"
       
   468 
       
   469 lemma minus_pinfreal_eq:
       
   470   "(x - y = (z :: pinfreal)) \<longleftrightarrow> (if y < x then x = y + z else z = 0)"
       
   471   (is "?diff \<longleftrightarrow> ?if")
       
   472 proof
       
   473   assume ?diff
       
   474   thus ?if
       
   475   proof (cases "y < x")
       
   476     case True
       
   477     then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
       
   478 
       
   479     show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pinfreal_def
       
   480     proof (rule theI2[where Q="\<lambda>d. x = y + d"])
       
   481       show "x = y + pinfreal_case (\<lambda>r. Real (r - real y)) \<omega> x" (is "x = y + ?d")
       
   482         using `y < x` p by (cases x) simp_all
       
   483 
       
   484       fix d assume "x = y + d"
       
   485       thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all
       
   486     qed simp
       
   487   qed (simp add: minus_pinfreal_def)
       
   488 next
       
   489   assume ?if
       
   490   thus ?diff
       
   491   proof (cases "y < x")
       
   492     case True
       
   493     then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
       
   494 
       
   495     from True `?if` have "x = y + z" by simp
       
   496 
       
   497     show ?thesis unfolding minus_pinfreal_def if_P[OF True] unfolding `x = y + z`
       
   498     proof (rule the_equality)
       
   499       fix d :: pinfreal assume "y + z = y + d"
       
   500       thus "d = z" using `y < x` p
       
   501         by (cases d, cases z) simp_all
       
   502     qed simp
       
   503   qed (simp add: minus_pinfreal_def)
       
   504 qed
       
   505 
       
   506 instance ..
       
   507 end
       
   508 
       
   509 lemma pinfreal_minus[simp]:
       
   510   "Real r - Real p = (if 0 \<le> r \<and> p < r then if 0 \<le> p then Real (r - p) else Real r else 0)"
       
   511   "(A::pinfreal) - A = 0"
       
   512   "\<omega> - Real r = \<omega>"
       
   513   "Real r - \<omega> = 0"
       
   514   "A - 0 = A"
       
   515   "0 - A = 0"
       
   516   by (auto simp: minus_pinfreal_eq not_less)
       
   517 
       
   518 lemma pinfreal_le_epsilon:
       
   519   fixes x y :: pinfreal
       
   520   assumes "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
       
   521   shows "x \<le> y"
       
   522 proof (cases y)
       
   523   case (preal r)
       
   524   then obtain p where x: "x = Real p" "0 \<le> p"
       
   525     using assms[of 1] by (cases x) auto
       
   526   { fix e have "0 < e \<Longrightarrow> p \<le> r + e"
       
   527       using assms[of "Real e"] preal x by auto }
       
   528   hence "p \<le> r" by (rule field_le_epsilon)
       
   529   thus ?thesis using preal x by auto
       
   530 qed simp
       
   531 
       
   532 instance pinfreal :: "{ordered_comm_semiring, comm_semiring_1}"
       
   533 proof
       
   534   show "0 \<noteq> (1::pinfreal)" unfolding zero_pinfreal_def one_pinfreal_def
       
   535     by (simp del: Real_1 Real_0)
       
   536 
       
   537   fix a :: pinfreal
       
   538   show "0 * a = 0" "a * 0 = 0" by simp_all
       
   539 
       
   540   fix b c :: pinfreal
       
   541   show "(a + b) * c = a * c + b * c"
       
   542     by (cases c, cases a, cases b)
       
   543        (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff)
       
   544 
       
   545   { assume "a \<le> b" thus "c + a \<le> c + b"
       
   546      by (cases c, cases a, cases b) auto }
       
   547 
       
   548   assume "a \<le> b" "0 \<le> c"
       
   549   thus "c * a \<le> c * b"
       
   550     apply (cases c, cases a, cases b)
       
   551     by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le)
       
   552 qed
       
   553 
       
   554 lemma mult_\<omega>[simp]: "x * y = \<omega> \<longleftrightarrow> (x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0"
       
   555   by (cases x, cases y) auto
       
   556 
       
   557 lemma \<omega>_mult[simp]: "(\<omega> = x * y) = ((x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0)"
       
   558   by (cases x, cases y) auto
       
   559 
       
   560 lemma pinfreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pinfreal) = 0"
       
   561   by (cases x, cases y) (auto simp: mult_le_0_iff)
       
   562 
       
   563 lemma pinfreal_mult_cancel:
       
   564   fixes x y z :: pinfreal
       
   565   assumes "y \<le> z"
       
   566   shows "x * y \<le> x * z"
       
   567   using assms
       
   568   by (cases x, cases y, cases z)
       
   569      (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le)
       
   570 
       
   571 lemma Real_power[simp]:
       
   572   "Real x ^ n = (if x \<le> 0 then (if n = 0 then 1 else 0) else Real (x ^ n))"
       
   573   by (induct n) auto
       
   574 
       
   575 lemma Real_power_\<omega>[simp]:
       
   576   "\<omega> ^ n = (if n = 0 then 1 else \<omega>)"
       
   577   by (induct n) auto
       
   578 
       
   579 lemma pinfreal_of_nat[simp]: "of_nat m = Real (real m)"
       
   580   by (induct m) (auto simp: real_of_nat_Suc one_pinfreal_def simp del: Real_1)
       
   581 
       
   582 lemma real_of_pinfreal_mono:
       
   583   fixes a b :: pinfreal
       
   584   assumes "b \<noteq> \<omega>" "a \<le> b"
       
   585   shows "real a \<le> real b"
       
   586 using assms by (cases b, cases a) auto
       
   587 
       
   588 instance pinfreal :: "semiring_char_0"
       
   589 proof
       
   590   fix m n
       
   591   show "inj (of_nat::nat\<Rightarrow>pinfreal)" by (auto intro!: inj_onI)
       
   592 qed
       
   593 
       
   594 subsection "@{typ pinfreal} is a complete lattice"
       
   595 
       
   596 instantiation pinfreal :: lattice
       
   597 begin
       
   598 definition [simp]: "sup x y = (max x y :: pinfreal)"
       
   599 definition [simp]: "inf x y = (min x y :: pinfreal)"
       
   600 instance proof qed simp_all
       
   601 end
       
   602 
       
   603 instantiation pinfreal :: complete_lattice
       
   604 begin
       
   605 
       
   606 definition "bot = Real 0"
       
   607 definition "top = \<omega>"
       
   608 
       
   609 definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pinfreal)"
       
   610 definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pinfreal)"
       
   611 
       
   612 lemma pinfreal_complete_Sup:
       
   613   fixes S :: "pinfreal set" assumes "S \<noteq> {}"
       
   614   shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
       
   615 proof (cases "\<exists>x\<ge>0. \<forall>a\<in>S. a \<le> Real x")
       
   616   case False
       
   617   hence *: "\<And>x. x\<ge>0 \<Longrightarrow> \<exists>a\<in>S. \<not>a \<le> Real x" by simp
       
   618   show ?thesis
       
   619   proof (safe intro!: exI[of _ \<omega>])
       
   620     fix y assume **: "\<forall>z\<in>S. z \<le> y"
       
   621     show "\<omega> \<le> y" unfolding pinfreal_\<omega>_less_eq
       
   622     proof (rule ccontr)
       
   623       assume "y \<noteq> \<omega>"
       
   624       then obtain x where [simp]: "y = Real x" and "0 \<le> x" by (cases y) auto
       
   625       from *[OF `0 \<le> x`] show False using ** by auto
       
   626     qed
       
   627   qed simp
       
   628 next
       
   629   case True then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> Real y" and "0 \<le> y" by auto
       
   630   from y[of \<omega>] have "\<omega> \<notin> S" by auto
       
   631 
       
   632   with `S \<noteq> {}` obtain x where "x \<in> S" and "x \<noteq> \<omega>" by auto
       
   633 
       
   634   have bound: "\<forall>x\<in>real ` S. x \<le> y"
       
   635   proof
       
   636     fix z assume "z \<in> real ` S" then guess a ..
       
   637     with y[of a] `\<omega> \<notin> S` `0 \<le> y` show "z \<le> y" by (cases a) auto
       
   638   qed
       
   639   with reals_complete2[of "real ` S"] `x \<in> S`
       
   640   obtain s where s: "\<forall>y\<in>S. real y \<le> s" "\<forall>z. ((\<forall>y\<in>S. real y \<le> z) \<longrightarrow> s \<le> z)"
       
   641     by auto
       
   642 
       
   643   show ?thesis
       
   644   proof (safe intro!: exI[of _ "Real s"])
       
   645     fix z assume "z \<in> S" thus "z \<le> Real s"
       
   646       using s `\<omega> \<notin> S` by (cases z) auto
       
   647   next
       
   648     fix z assume *: "\<forall>y\<in>S. y \<le> z"
       
   649     show "Real s \<le> z"
       
   650     proof (cases z)
       
   651       case (preal u)
       
   652       { fix v assume "v \<in> S"
       
   653         hence "v \<le> Real u" using * preal by auto
       
   654         hence "real v \<le> u" using `\<omega> \<notin> S` `0 \<le> u` by (cases v) auto }
       
   655       hence "s \<le> u" using s(2) by auto
       
   656       thus "Real s \<le> z" using preal by simp
       
   657     qed simp
       
   658   qed
       
   659 qed
       
   660 
       
   661 lemma pinfreal_complete_Inf:
       
   662   fixes S :: "pinfreal set" assumes "S \<noteq> {}"
       
   663   shows "\<exists>x. (\<forall>y\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> x)"
       
   664 proof (cases "S = {\<omega>}")
       
   665   case True thus ?thesis by (auto intro!: exI[of _ \<omega>])
       
   666 next
       
   667   case False with `S \<noteq> {}` have "S - {\<omega>} \<noteq> {}" by auto
       
   668   hence not_empty: "\<exists>x. x \<in> uminus ` real ` (S - {\<omega>})" by auto
       
   669   have bounds: "\<exists>x. \<forall>y\<in>uminus ` real ` (S - {\<omega>}). y \<le> x" by (auto intro!: exI[of _ 0])
       
   670   from reals_complete2[OF not_empty bounds]
       
   671   obtain s where s: "\<And>y. y\<in>S - {\<omega>} \<Longrightarrow> - real y \<le> s" "\<forall>z. ((\<forall>y\<in>S - {\<omega>}. - real y \<le> z) \<longrightarrow> s \<le> z)"
       
   672     by auto
       
   673 
       
   674   show ?thesis
       
   675   proof (safe intro!: exI[of _ "Real (-s)"])
       
   676     fix z assume "z \<in> S"
       
   677     show "Real (-s) \<le> z"
       
   678     proof (cases z)
       
   679       case (preal r)
       
   680       with s `z \<in> S` have "z \<in> S - {\<omega>}" by simp
       
   681       hence "- r \<le> s" using preal s(1)[of z] by auto
       
   682       hence "- s \<le> r" by (subst neg_le_iff_le[symmetric]) simp
       
   683       thus ?thesis using preal by simp
       
   684     qed simp
       
   685   next
       
   686     fix z assume *: "\<forall>y\<in>S. z \<le> y"
       
   687     show "z \<le> Real (-s)"
       
   688     proof (cases z)
       
   689       case (preal u)
       
   690       { fix v assume "v \<in> S-{\<omega>}"
       
   691         hence "Real u \<le> v" using * preal by auto
       
   692         hence "- real v \<le> - u" using `0 \<le> u` `v \<in> S - {\<omega>}` by (cases v) auto }
       
   693       hence "u \<le> - s" using s(2) by (subst neg_le_iff_le[symmetric]) auto
       
   694       thus "z \<le> Real (-s)" using preal by simp
       
   695     next
       
   696       case infinite
       
   697       with * have "S = {\<omega>}" using `S \<noteq> {}` by auto
       
   698       with `S - {\<omega>} \<noteq> {}` show ?thesis by simp
       
   699     qed
       
   700   qed
       
   701 qed
       
   702 
       
   703 instance
       
   704 proof
       
   705   fix x :: pinfreal and A
       
   706 
       
   707   show "bot \<le> x" by (cases x) (simp_all add: bot_pinfreal_def)
       
   708   show "x \<le> top" by (simp add: top_pinfreal_def)
       
   709 
       
   710   { assume "x \<in> A"
       
   711     with pinfreal_complete_Sup[of A]
       
   712     obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
       
   713     hence "x \<le> s" using `x \<in> A` by auto
       
   714     also have "... = Sup A" using s unfolding Sup_pinfreal_def
       
   715       by (auto intro!: Least_equality[symmetric])
       
   716     finally show "x \<le> Sup A" . }
       
   717 
       
   718   { assume "x \<in> A"
       
   719     with pinfreal_complete_Inf[of A]
       
   720     obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
       
   721     hence "Inf A = i" unfolding Inf_pinfreal_def
       
   722       by (auto intro!: Greatest_equality)
       
   723     also have "i \<le> x" using i `x \<in> A` by auto
       
   724     finally show "Inf A \<le> x" . }
       
   725 
       
   726   { assume *: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x"
       
   727     show "Sup A \<le> x"
       
   728     proof (cases "A = {}")
       
   729       case True
       
   730       hence "Sup A = 0" unfolding Sup_pinfreal_def
       
   731         by (auto intro!: Least_equality)
       
   732       thus "Sup A \<le> x" by simp
       
   733     next
       
   734       case False
       
   735       with pinfreal_complete_Sup[of A]
       
   736       obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
       
   737       hence "Sup A = s"
       
   738         unfolding Sup_pinfreal_def by (auto intro!: Least_equality)
       
   739       also have "s \<le> x" using * s by auto
       
   740       finally show "Sup A \<le> x" .
       
   741     qed }
       
   742 
       
   743   { assume *: "\<And>z. z \<in> A \<Longrightarrow> x \<le> z"
       
   744     show "x \<le> Inf A"
       
   745     proof (cases "A = {}")
       
   746       case True
       
   747       hence "Inf A = \<omega>" unfolding Inf_pinfreal_def
       
   748         by (auto intro!: Greatest_equality)
       
   749       thus "x \<le> Inf A" by simp
       
   750     next
       
   751       case False
       
   752       with pinfreal_complete_Inf[of A]
       
   753       obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
       
   754       have "x \<le> i" using * i by auto
       
   755       also have "i = Inf A" using i
       
   756         unfolding Inf_pinfreal_def by (auto intro!: Greatest_equality[symmetric])
       
   757       finally show "x \<le> Inf A" .
       
   758     qed }
       
   759 qed
       
   760 end
       
   761 
       
   762 lemma Inf_pinfreal_iff:
       
   763   fixes z :: pinfreal
       
   764   shows "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> (\<exists>x\<in>X. x<y) \<longleftrightarrow> Inf X < y"
       
   765   by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
       
   766             order_less_le_trans)
       
   767 
       
   768 lemma Inf_greater:
       
   769   fixes z :: pinfreal assumes "Inf X < z"
       
   770   shows "\<exists>x \<in> X. x < z"
       
   771 proof -
       
   772   have "X \<noteq> {}" using assms by (auto simp: Inf_empty top_pinfreal_def)
       
   773   with assms show ?thesis
       
   774     by (metis Inf_pinfreal_iff mem_def not_leE)
       
   775 qed
       
   776 
       
   777 lemma Inf_close:
       
   778   fixes e :: pinfreal assumes "Inf X \<noteq> \<omega>" "0 < e"
       
   779   shows "\<exists>x \<in> X. x < Inf X + e"
       
   780 proof (rule Inf_greater)
       
   781   show "Inf X < Inf X + e" using assms
       
   782     by (cases "Inf X", cases e) auto
       
   783 qed
       
   784 
       
   785 lemma pinfreal_SUPI:
       
   786   fixes x :: pinfreal
       
   787   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
       
   788   assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y"
       
   789   shows "(SUP i:A. f i) = x"
       
   790   unfolding SUPR_def Sup_pinfreal_def
       
   791   using assms by (auto intro!: Least_equality)
       
   792 
       
   793 lemma Sup_pinfreal_iff:
       
   794   fixes z :: pinfreal
       
   795   shows "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> (\<exists>x\<in>X. y<x) \<longleftrightarrow> y < Sup X"
       
   796   by (metis complete_lattice_class.Sup_least complete_lattice_class.Sup_upper less_le_not_le linear
       
   797             order_less_le_trans)
       
   798 
       
   799 lemma Sup_lesser:
       
   800   fixes z :: pinfreal assumes "z < Sup X"
       
   801   shows "\<exists>x \<in> X. z < x"
       
   802 proof -
       
   803   have "X \<noteq> {}" using assms by (auto simp: Sup_empty bot_pinfreal_def)
       
   804   with assms show ?thesis
       
   805     by (metis Sup_pinfreal_iff mem_def not_leE)
       
   806 qed
       
   807 
       
   808 lemma Sup_eq_\<omega>: "\<omega> \<in> S \<Longrightarrow> Sup S = \<omega>"
       
   809   unfolding Sup_pinfreal_def
       
   810   by (auto intro!: Least_equality)
       
   811 
       
   812 lemma Sup_close:
       
   813   assumes "0 < e" and S: "Sup S \<noteq> \<omega>" "S \<noteq> {}"
       
   814   shows "\<exists>X\<in>S. Sup S < X + e"
       
   815 proof cases
       
   816   assume "Sup S = 0"
       
   817   moreover obtain X where "X \<in> S" using `S \<noteq> {}` by auto
       
   818   ultimately show ?thesis using `0 < e` by (auto intro!: bexI[OF _ `X\<in>S`])
       
   819 next
       
   820   assume "Sup S \<noteq> 0"
       
   821   have "\<exists>X\<in>S. Sup S - e < X"
       
   822   proof (rule Sup_lesser)
       
   823     show "Sup S - e < Sup S" using `0 < e` `Sup S \<noteq> 0` `Sup S \<noteq> \<omega>`
       
   824       by (cases e) (auto simp: pinfreal_noteq_omega_Ex)
       
   825   qed
       
   826   then guess X .. note X = this
       
   827   with `Sup S \<noteq> \<omega>` Sup_eq_\<omega> have "X \<noteq> \<omega>" by auto
       
   828   thus ?thesis using `Sup S \<noteq> \<omega>` X unfolding pinfreal_noteq_omega_Ex
       
   829     by (cases e) (auto intro!: bexI[OF _ `X\<in>S`] simp: split: split_if_asm)
       
   830 qed
       
   831 
       
   832 lemma Sup_\<omega>: "(SUP i::nat. Real (real i)) = \<omega>"
       
   833 proof (rule pinfreal_SUPI)
       
   834   fix y assume *: "\<And>i::nat. i \<in> UNIV \<Longrightarrow> Real (real i) \<le> y"
       
   835   thus "\<omega> \<le> y"
       
   836   proof (cases y)
       
   837     case (preal r)
       
   838     then obtain k :: nat where "r < real k"
       
   839       using ex_less_of_nat by (auto simp: real_eq_of_nat)
       
   840     with *[of k] preal show ?thesis by auto
       
   841   qed simp
       
   842 qed simp
       
   843 
       
   844 subsubsection {* Equivalence between @{text "f ----> x"} and @{text SUP} on @{typ pinfreal} *}
       
   845 
       
   846 lemma monoseq_monoI: "mono f \<Longrightarrow> monoseq f"
       
   847   unfolding mono_def monoseq_def by auto
       
   848 
       
   849 lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
       
   850   unfolding mono_def incseq_def by auto
       
   851 
       
   852 lemma SUP_eq_LIMSEQ:
       
   853   assumes "mono f" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
       
   854   shows "(SUP n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
       
   855 proof
       
   856   assume x: "(SUP n. Real (f n)) = Real x"
       
   857   { fix n
       
   858     have "Real (f n) \<le> Real x" using x[symmetric] by (auto intro: le_SUPI)
       
   859     hence "f n \<le> x" using assms by simp }
       
   860   show "f ----> x"
       
   861   proof (rule LIMSEQ_I)
       
   862     fix r :: real assume "0 < r"
       
   863     show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
       
   864     proof (rule ccontr)
       
   865       assume *: "\<not> ?thesis"
       
   866       { fix N
       
   867         from * obtain n where "N \<le> n" "r \<le> x - f n"
       
   868           using `\<And>n. f n \<le> x` by (auto simp: not_less)
       
   869         hence "f N \<le> f n" using `mono f` by (auto dest: monoD)
       
   870         hence "f N \<le> x - r" using `r \<le> x - f n` by auto
       
   871         hence "Real (f N) \<le> Real (x - r)" and "r \<le> x" using `0 \<le> f N` by auto }
       
   872       hence "(SUP n. Real (f n)) \<le> Real (x - r)"
       
   873         and "Real (x - r) < Real x" using `0 < r` by (auto intro: SUP_leI)
       
   874       hence "(SUP n. Real (f n)) < Real x" by (rule le_less_trans)
       
   875       thus False using x by auto
       
   876     qed
       
   877   qed
       
   878 next
       
   879   assume "f ----> x"
       
   880   show "(SUP n. Real (f n)) = Real x"
       
   881   proof (rule pinfreal_SUPI)
       
   882     fix n
       
   883     from incseq_le[of f x] `mono f` `f ----> x`
       
   884     show "Real (f n) \<le> Real x" using assms incseq_mono by auto
       
   885   next
       
   886     fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> Real (f n) \<le> y"
       
   887     show "Real x \<le> y"
       
   888     proof (cases y)
       
   889       case (preal r)
       
   890       with * have "\<exists>N. \<forall>n\<ge>N. f n \<le> r" using assms by fastsimp
       
   891       from LIMSEQ_le_const2[OF `f ----> x` this]
       
   892       show "Real x \<le> y" using `0 \<le> x` preal by auto
       
   893     qed simp
       
   894   qed
       
   895 qed
       
   896 
       
   897 lemma SUPR_bound:
       
   898   assumes "\<forall>N. f N \<le> x"
       
   899   shows "(SUP n. f n) \<le> x"
       
   900   using assms by (simp add: SUPR_def Sup_le_iff)
       
   901 
       
   902 lemma pinfreal_less_eq_diff_eq_sum:
       
   903   fixes x y z :: pinfreal
       
   904   assumes "y \<le> x" and "x \<noteq> \<omega>"
       
   905   shows "z \<le> x - y \<longleftrightarrow> z + y \<le> x"
       
   906   using assms
       
   907   apply (cases z, cases y, cases x)
       
   908   by (simp_all add: field_simps minus_pinfreal_eq)
       
   909 
       
   910 lemma Real_diff_less_omega: "Real r - x < \<omega>" by (cases x) auto
       
   911 
       
   912 subsubsection {* Numbers on @{typ pinfreal} *}
       
   913 
       
   914 instantiation pinfreal :: number
       
   915 begin
       
   916 definition [simp]: "number_of x = Real (number_of x)"
       
   917 instance proof qed
       
   918 end
       
   919 
       
   920 subsubsection {* Division on @{typ pinfreal} *}
       
   921 
       
   922 instantiation pinfreal :: inverse
       
   923 begin
       
   924 
       
   925 definition "inverse x = pinfreal_case (\<lambda>x. if x = 0 then \<omega> else Real (inverse x)) 0 x"
       
   926 definition [simp]: "x / y = x * inverse (y :: pinfreal)"
       
   927 
       
   928 instance proof qed
       
   929 end
       
   930 
       
   931 lemma pinfreal_inverse[simp]:
       
   932   "inverse 0 = \<omega>"
       
   933   "inverse (Real x) = (if x \<le> 0 then \<omega> else Real (inverse x))"
       
   934   "inverse \<omega> = 0"
       
   935   "inverse (1::pinfreal) = 1"
       
   936   "inverse (inverse x) = x"
       
   937   by (simp_all add: inverse_pinfreal_def one_pinfreal_def split: pinfreal_case_split del: Real_1)
       
   938 
       
   939 lemma pinfreal_inverse_le_eq:
       
   940   assumes "x \<noteq> 0" "x \<noteq> \<omega>"
       
   941   shows "y \<le> z / x \<longleftrightarrow> x * y \<le> (z :: pinfreal)"
       
   942 proof -
       
   943   from assms obtain r where r: "x = Real r" "0 < r" by (cases x) auto
       
   944   { fix p q :: real assume "0 \<le> p" "0 \<le> q"
       
   945     have "p \<le> q * inverse r \<longleftrightarrow> p \<le> q / r" by (simp add: divide_inverse)
       
   946     also have "... \<longleftrightarrow> p * r \<le> q" using `0 < r` by (auto simp: field_simps)
       
   947     finally have "p \<le> q * inverse r \<longleftrightarrow> p * r \<le> q" . }
       
   948   with r show ?thesis
       
   949     by (cases y, cases z, auto simp: zero_le_mult_iff field_simps)
       
   950 qed
       
   951 
       
   952 lemma inverse_antimono_strict:
       
   953   fixes x y :: pinfreal
       
   954   assumes "x < y" shows "inverse y < inverse x"
       
   955   using assms by (cases x, cases y) auto
       
   956 
       
   957 lemma inverse_antimono:
       
   958   fixes x y :: pinfreal
       
   959   assumes "x \<le> y" shows "inverse y \<le> inverse x"
       
   960   using assms by (cases x, cases y) auto
       
   961 
       
   962 lemma pinfreal_inverse_\<omega>_iff[simp]: "inverse x = \<omega> \<longleftrightarrow> x = 0"
       
   963   by (cases x) auto
       
   964 
       
   965 subsection "Infinite sum over @{typ pinfreal}"
       
   966 
       
   967 text {*
       
   968 
       
   969 The infinite sum over @{typ pinfreal} has the nice property that it is always
       
   970 defined.
       
   971 
       
   972 *}
       
   973 
       
   974 definition psuminf :: "(nat \<Rightarrow> pinfreal) \<Rightarrow> pinfreal" (binder "\<Sum>\<^isub>\<infinity>" 10) where
       
   975   "(\<Sum>\<^isub>\<infinity> x. f x) = (SUP n. \<Sum>i<n. f i)"
       
   976 
       
   977 subsubsection {* Equivalence between @{text "\<Sum> n. f n"} and @{text "\<Sum>\<^isub>\<infinity> n. f n"} *}
       
   978 
       
   979 lemma setsum_Real:
       
   980   assumes "\<forall>x\<in>A. 0 \<le> f x"
       
   981   shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
       
   982 proof (cases "finite A")
       
   983   case True
       
   984   thus ?thesis using assms
       
   985   proof induct case (insert x s)
       
   986     hence "0 \<le> setsum f s" apply-apply(rule setsum_nonneg) by auto
       
   987     thus ?case using insert by auto
       
   988   qed auto
       
   989 qed simp
       
   990 
       
   991 lemma setsum_Real':
       
   992   assumes "\<forall>x. 0 \<le> f x"
       
   993   shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
       
   994   apply(rule setsum_Real) using assms by auto
       
   995 
       
   996 lemma setsum_\<omega>:
       
   997   "(\<Sum>x\<in>P. f x) = \<omega> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<omega>))"
       
   998 proof safe
       
   999   assume *: "setsum f P = \<omega>"
       
  1000   show "finite P"
       
  1001   proof (rule ccontr) assume "infinite P" with * show False by auto qed
       
  1002   show "\<exists>i\<in>P. f i = \<omega>"
       
  1003   proof (rule ccontr)
       
  1004     assume "\<not> ?thesis" hence "\<And>i. i\<in>P \<Longrightarrow> f i \<noteq> \<omega>" by auto
       
  1005     from `finite P` this have "setsum f P \<noteq> \<omega>"
       
  1006       by induct auto
       
  1007     with * show False by auto
       
  1008   qed
       
  1009 next
       
  1010   fix i assume "finite P" "i \<in> P" "f i = \<omega>"
       
  1011   thus "setsum f P = \<omega>"
       
  1012   proof induct
       
  1013     case (insert x A)
       
  1014     show ?case using insert by (cases "x = i") auto
       
  1015   qed simp
       
  1016 qed
       
  1017 
       
  1018 lemma real_of_pinfreal_setsum:
       
  1019   assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> \<omega>"
       
  1020   shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
       
  1021 proof cases
       
  1022   assume "finite S"
       
  1023   from this assms show ?thesis
       
  1024     by induct (simp_all add: real_of_pinfreal_add setsum_\<omega>)
       
  1025 qed simp
       
  1026 
       
  1027 lemma setsum_0:
       
  1028   fixes f :: "'a \<Rightarrow> pinfreal" assumes "finite A"
       
  1029   shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
       
  1030   using assms by induct auto
       
  1031 
       
  1032 lemma suminf_imp_psuminf:
       
  1033   assumes "f sums x" and "\<forall>n. 0 \<le> f n"
       
  1034   shows "(\<Sum>\<^isub>\<infinity> x. Real (f x)) = Real x"
       
  1035   unfolding psuminf_def setsum_Real'[OF assms(2)]
       
  1036 proof (rule SUP_eq_LIMSEQ[THEN iffD2])
       
  1037   show "mono (\<lambda>n. setsum f {..<n})" (is "mono ?S")
       
  1038     unfolding mono_iff_le_Suc using assms by simp
       
  1039 
       
  1040   { fix n show "0 \<le> ?S n"
       
  1041       using setsum_nonneg[of "{..<n}" f] assms by auto }
       
  1042 
       
  1043   thus "0 \<le> x" "?S ----> x"
       
  1044     using `f sums x` LIMSEQ_le_const
       
  1045     by (auto simp: atLeast0LessThan sums_def)
       
  1046 qed
       
  1047 
       
  1048 lemma psuminf_equality:
       
  1049   assumes "\<And>n. setsum f {..<n} \<le> x"
       
  1050   and "\<And>y. y \<noteq> \<omega> \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> y) \<Longrightarrow> x \<le> y"
       
  1051   shows "psuminf f = x"
       
  1052   unfolding psuminf_def
       
  1053 proof (safe intro!: pinfreal_SUPI)
       
  1054   fix n show "setsum f {..<n} \<le> x" using assms(1) .
       
  1055 next
       
  1056   fix y assume *: "\<forall>n. n \<in> UNIV \<longrightarrow> setsum f {..<n} \<le> y"
       
  1057   show "x \<le> y"
       
  1058   proof (cases "y = \<omega>")
       
  1059     assume "y \<noteq> \<omega>" from assms(2)[OF this] *
       
  1060     show "x \<le> y" by auto
       
  1061   qed simp
       
  1062 qed
       
  1063 
       
  1064 lemma psuminf_\<omega>:
       
  1065   assumes "f i = \<omega>"
       
  1066   shows "(\<Sum>\<^isub>\<infinity> x. f x) = \<omega>"
       
  1067 proof (rule psuminf_equality)
       
  1068   fix y assume *: "\<And>n. setsum f {..<n} \<le> y"
       
  1069   have "setsum f {..<Suc i} = \<omega>" 
       
  1070     using assms by (simp add: setsum_\<omega>)
       
  1071   thus "\<omega> \<le> y" using *[of "Suc i"] by auto
       
  1072 qed simp
       
  1073 
       
  1074 lemma psuminf_imp_suminf:
       
  1075   assumes "(\<Sum>\<^isub>\<infinity> x. f x) \<noteq> \<omega>"
       
  1076   shows "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity> x. f x)"
       
  1077 proof -
       
  1078   have "\<forall>i. \<exists>r. f i = Real r \<and> 0 \<le> r"
       
  1079   proof
       
  1080     fix i show "\<exists>r. f i = Real r \<and> 0 \<le> r" using psuminf_\<omega> assms by (cases "f i") auto
       
  1081   qed
       
  1082   from choice[OF this] obtain r where f: "f = (\<lambda>i. Real (r i))"
       
  1083     and pos: "\<forall>i. 0 \<le> r i"
       
  1084     by (auto simp: expand_fun_eq)
       
  1085   hence [simp]: "\<And>i. real (f i) = r i" by auto
       
  1086 
       
  1087   have "mono (\<lambda>n. setsum r {..<n})" (is "mono ?S")
       
  1088     unfolding mono_iff_le_Suc using pos by simp
       
  1089 
       
  1090   { fix n have "0 \<le> ?S n"
       
  1091       using setsum_nonneg[of "{..<n}" r] pos by auto }
       
  1092 
       
  1093   from assms obtain p where *: "(\<Sum>\<^isub>\<infinity> x. f x) = Real p" and "0 \<le> p"
       
  1094     by (cases "(\<Sum>\<^isub>\<infinity> x. f x)") auto
       
  1095   show ?thesis unfolding * using * pos `0 \<le> p` SUP_eq_LIMSEQ[OF `mono ?S` `\<And>n. 0 \<le> ?S n` `0 \<le> p`]
       
  1096     by (simp add: f atLeast0LessThan sums_def psuminf_def setsum_Real'[OF pos] f)
       
  1097 qed
       
  1098 
       
  1099 lemma psuminf_bound:
       
  1100   assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x"
       
  1101   shows "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x"
       
  1102   using assms by (simp add: psuminf_def SUPR_def Sup_le_iff)
       
  1103 
       
  1104 lemma psuminf_bound_add:
       
  1105   assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x"
       
  1106   shows "(\<Sum>\<^isub>\<infinity> n. f n) + y \<le> x"
       
  1107 proof (cases "x = \<omega>")
       
  1108   have "y \<le> x" using assms by (auto intro: pinfreal_le_add2)
       
  1109   assume "x \<noteq> \<omega>"
       
  1110   note move_y = pinfreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
       
  1111 
       
  1112   have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" using assms by (simp add: move_y)
       
  1113   hence "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x - y" by (rule psuminf_bound)
       
  1114   thus ?thesis by (simp add: move_y)
       
  1115 qed simp
       
  1116 
       
  1117 lemma psuminf_finite:
       
  1118   assumes "\<forall>N\<ge>n. f N = 0"
       
  1119   shows "(\<Sum>\<^isub>\<infinity> n. f n) = (\<Sum>N<n. f N)"
       
  1120 proof (rule psuminf_equality)
       
  1121   fix N
       
  1122   show "setsum f {..<N} \<le> setsum f {..<n}"
       
  1123   proof (cases rule: linorder_cases)
       
  1124     assume "N < n" thus ?thesis by (auto intro!: setsum_mono3)
       
  1125   next
       
  1126     assume "n < N"
       
  1127     hence *: "{..<N} = {..<n} \<union> {n..<N}" by auto
       
  1128     moreover have "setsum f {n..<N} = 0"
       
  1129       using assms by (auto intro!: setsum_0')
       
  1130     ultimately show ?thesis unfolding *
       
  1131       by (subst setsum_Un_disjoint) auto
       
  1132   qed simp
       
  1133 qed simp
       
  1134 
       
  1135 lemma psuminf_upper:
       
  1136   shows "(\<Sum>n<N. f n) \<le> (\<Sum>\<^isub>\<infinity> n. f n)"
       
  1137   unfolding psuminf_def SUPR_def
       
  1138   by (auto intro: complete_lattice_class.Sup_upper image_eqI)
       
  1139 
       
  1140 lemma psuminf_le:
       
  1141   assumes "\<And>N. f N \<le> g N"
       
  1142   shows "psuminf f \<le> psuminf g"
       
  1143 proof (safe intro!: psuminf_bound)
       
  1144   fix n
       
  1145   have "setsum f {..<n} \<le> setsum g {..<n}" using assms by (auto intro: setsum_mono)
       
  1146   also have "... \<le> psuminf g" by (rule psuminf_upper)
       
  1147   finally show "setsum f {..<n} \<le> psuminf g" .
       
  1148 qed
       
  1149 
       
  1150 lemma psuminf_const[simp]: "psuminf (\<lambda>n. c) = (if c = 0 then 0 else \<omega>)" (is "_ = ?if")
       
  1151 proof (rule psuminf_equality)
       
  1152   fix y assume *: "\<And>n :: nat. (\<Sum>n<n. c) \<le> y" and "y \<noteq> \<omega>"
       
  1153   then obtain r p where
       
  1154     y: "y = Real r" "0 \<le> r" and
       
  1155     c: "c = Real p" "0 \<le> p"
       
  1156       using *[of 1] by (cases c, cases y) auto
       
  1157   show "(if c = 0 then 0 else \<omega>) \<le> y"
       
  1158   proof (cases "p = 0")
       
  1159     assume "p = 0" with c show ?thesis by simp
       
  1160   next
       
  1161     assume "p \<noteq> 0"
       
  1162     with * c y have **: "\<And>n :: nat. real n \<le> r / p"
       
  1163       by (auto simp: zero_le_mult_iff field_simps)
       
  1164     from ex_less_of_nat[of "r / p"] guess n ..
       
  1165     with **[of n] show ?thesis by (simp add: real_eq_of_nat)
       
  1166   qed
       
  1167 qed (cases "c = 0", simp_all)
       
  1168 
       
  1169 lemma psuminf_add[simp]: "psuminf (\<lambda>n. f n + g n) = psuminf f + psuminf g"
       
  1170 proof (rule psuminf_equality)
       
  1171   fix n
       
  1172   from psuminf_upper[of f n] psuminf_upper[of g n]
       
  1173   show "(\<Sum>n<n. f n + g n) \<le> psuminf f + psuminf g"
       
  1174     by (auto simp add: setsum_addf intro!: add_mono)
       
  1175 next
       
  1176   fix y assume *: "\<And>n. (\<Sum>n<n. f n + g n) \<le> y" and "y \<noteq> \<omega>"
       
  1177   { fix n m
       
  1178     have **: "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y"
       
  1179     proof (cases rule: linorder_le_cases)
       
  1180       assume "n \<le> m"
       
  1181       hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<m. f n) + (\<Sum>n<m. g n)"
       
  1182         by (auto intro!: add_right_mono setsum_mono3)
       
  1183       also have "... \<le> y"
       
  1184         using * by (simp add: setsum_addf)
       
  1185       finally show ?thesis .
       
  1186     next
       
  1187       assume "m \<le> n"
       
  1188       hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<n. f n) + (\<Sum>n<n. g n)"
       
  1189         by (auto intro!: add_left_mono setsum_mono3)
       
  1190       also have "... \<le> y"
       
  1191         using * by (simp add: setsum_addf)
       
  1192       finally show ?thesis .
       
  1193     qed }
       
  1194   hence "\<And>m. \<forall>n. (\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y" by simp
       
  1195   from psuminf_bound_add[OF this]
       
  1196   have "\<forall>m. (\<Sum>n<m. g n) + psuminf f \<le> y" by (simp add: ac_simps)
       
  1197   from psuminf_bound_add[OF this]
       
  1198   show "psuminf f + psuminf g \<le> y" by (simp add: ac_simps)
       
  1199 qed
       
  1200 
       
  1201 lemma psuminf_0: "psuminf f = 0 \<longleftrightarrow> (\<forall>i. f i = 0)"
       
  1202 proof safe
       
  1203   assume "\<forall>i. f i = 0"
       
  1204   hence "f = (\<lambda>i. 0)" by (simp add: expand_fun_eq)
       
  1205   thus "psuminf f = 0" using psuminf_const by simp
       
  1206 next
       
  1207   fix i assume "psuminf f = 0"
       
  1208   hence "(\<Sum>n<Suc i. f n) = 0" using psuminf_upper[of f "Suc i"] by simp
       
  1209   thus "f i = 0" by simp
       
  1210 qed
       
  1211 
       
  1212 lemma psuminf_cmult_right[simp]: "psuminf (\<lambda>n. c * f n) = c * psuminf f"
       
  1213 proof (rule psuminf_equality)
       
  1214   fix n show "(\<Sum>n<n. c * f n) \<le> c * psuminf f"
       
  1215    by (auto simp: setsum_right_distrib[symmetric] intro: mult_left_mono psuminf_upper)
       
  1216 next
       
  1217   fix y
       
  1218   assume "\<And>n. (\<Sum>n<n. c * f n) \<le> y"
       
  1219   hence *: "\<And>n. c * (\<Sum>n<n. f n) \<le> y" by (auto simp add: setsum_right_distrib)
       
  1220   thus "c * psuminf f \<le> y"
       
  1221   proof (cases "c = \<omega> \<or> c = 0")
       
  1222     assume "c = \<omega> \<or> c = 0"
       
  1223     thus ?thesis
       
  1224       using * by (fastsimp simp add: psuminf_0 setsum_0 split: split_if_asm)
       
  1225   next
       
  1226     assume "\<not> (c = \<omega> \<or> c = 0)"
       
  1227     hence "c \<noteq> 0" "c \<noteq> \<omega>" by auto
       
  1228     note rewrite_div = pinfreal_inverse_le_eq[OF this, of _ y]
       
  1229     hence "\<forall>n. (\<Sum>n<n. f n) \<le> y / c" using * by simp
       
  1230     hence "psuminf f \<le> y / c" by (rule psuminf_bound)
       
  1231     thus ?thesis using rewrite_div by simp
       
  1232   qed
       
  1233 qed
       
  1234 
       
  1235 lemma psuminf_cmult_left[simp]: "psuminf (\<lambda>n. f n * c) = psuminf f * c"
       
  1236   using psuminf_cmult_right[of c f] by (simp add: ac_simps)
       
  1237 
       
  1238 lemma psuminf_half_series: "psuminf (\<lambda>n. (1/2)^Suc n) = 1"
       
  1239   using suminf_imp_psuminf[OF power_half_series] by auto
       
  1240 
       
  1241 lemma setsum_pinfsum: "(\<Sum>\<^isub>\<infinity> n. \<Sum>m\<in>A. f n m) = (\<Sum>m\<in>A. (\<Sum>\<^isub>\<infinity> n. f n m))"
       
  1242 proof (cases "finite A")
       
  1243   assume "finite A"
       
  1244   thus ?thesis by induct simp_all
       
  1245 qed simp
       
  1246 
       
  1247 lemma psuminf_reindex:
       
  1248   fixes f:: "nat \<Rightarrow> nat" assumes "bij f"
       
  1249   shows "psuminf (g \<circ> f) = psuminf g"
       
  1250 proof -
       
  1251   have [intro, simp]: "\<And>A. inj_on f A" using `bij f` unfolding bij_def by (auto intro: subset_inj_on)
       
  1252   have f[intro, simp]: "\<And>x. f (inv f x) = x"
       
  1253     using `bij f` unfolding bij_def by (auto intro: surj_f_inv_f)
       
  1254 
       
  1255   show ?thesis
       
  1256   proof (rule psuminf_equality)
       
  1257     fix n
       
  1258     have "setsum (g \<circ> f) {..<n} = setsum g (f ` {..<n})"
       
  1259       by (simp add: setsum_reindex)
       
  1260     also have "\<dots> \<le> setsum g {..Max (f ` {..<n})}"
       
  1261       by (rule setsum_mono3) auto
       
  1262     also have "\<dots> \<le> psuminf g" unfolding lessThan_Suc_atMost[symmetric] by (rule psuminf_upper)
       
  1263     finally show "setsum (g \<circ> f) {..<n} \<le> psuminf g" .
       
  1264   next
       
  1265     fix y assume *: "\<And>n. setsum (g \<circ> f) {..<n} \<le> y"
       
  1266     show "psuminf g \<le> y"
       
  1267     proof (safe intro!: psuminf_bound)
       
  1268       fix N
       
  1269       have "setsum g {..<N} \<le> setsum g (f ` {..Max (inv f ` {..<N})})"
       
  1270         by (rule setsum_mono3) (auto intro!: image_eqI[where f="f", OF f[symmetric]])
       
  1271       also have "\<dots> = setsum (g \<circ> f) {..Max (inv f ` {..<N})}"
       
  1272         by (simp add: setsum_reindex)
       
  1273       also have "\<dots> \<le> y" unfolding lessThan_Suc_atMost[symmetric] by (rule *)
       
  1274       finally show "setsum g {..<N} \<le> y" .
       
  1275     qed
       
  1276   qed
       
  1277 qed
       
  1278 
       
  1279 lemma psuminf_2dimen:
       
  1280   fixes f:: "nat * nat \<Rightarrow> pinfreal"
       
  1281   assumes fsums: "\<And>m. g m = (\<Sum>\<^isub>\<infinity> n. f (m,n))"
       
  1282   shows "psuminf (f \<circ> prod_decode) = psuminf g"
       
  1283 proof (rule psuminf_equality)
       
  1284   fix n :: nat
       
  1285   let ?P = "prod_decode ` {..<n}"
       
  1286   have "setsum (f \<circ> prod_decode) {..<n} = setsum f ?P"
       
  1287     by (auto simp: setsum_reindex inj_prod_decode)
       
  1288   also have "\<dots> \<le> setsum f ({..Max (fst ` ?P)} \<times> {..Max (snd ` ?P)})"
       
  1289   proof (safe intro!: setsum_mono3 Max_ge image_eqI)
       
  1290     fix a b x assume "(a, b) = prod_decode x"
       
  1291     from this[symmetric] show "a = fst (prod_decode x)" "b = snd (prod_decode x)"
       
  1292       by simp_all
       
  1293   qed simp_all
       
  1294   also have "\<dots> = (\<Sum>m\<le>Max (fst ` ?P). (\<Sum>n\<le>Max (snd ` ?P). f (m,n)))"
       
  1295     unfolding setsum_cartesian_product by simp
       
  1296   also have "\<dots> \<le> (\<Sum>m\<le>Max (fst ` ?P). g m)"
       
  1297     by (auto intro!: setsum_mono psuminf_upper simp del: setsum_lessThan_Suc
       
  1298         simp: fsums lessThan_Suc_atMost[symmetric])
       
  1299   also have "\<dots> \<le> psuminf g"
       
  1300     by (auto intro!: psuminf_upper simp del: setsum_lessThan_Suc
       
  1301         simp: lessThan_Suc_atMost[symmetric])
       
  1302   finally show "setsum (f \<circ> prod_decode) {..<n} \<le> psuminf g" .
       
  1303 next
       
  1304   fix y assume *: "\<And>n. setsum (f \<circ> prod_decode) {..<n} \<le> y"
       
  1305   have g: "g = (\<lambda>m. \<Sum>\<^isub>\<infinity> n. f (m,n))" unfolding fsums[symmetric] ..
       
  1306   show "psuminf g \<le> y" unfolding g
       
  1307   proof (rule psuminf_bound, unfold setsum_pinfsum[symmetric], safe intro!: psuminf_bound)
       
  1308     fix N M :: nat
       
  1309     let ?P = "{..<N} \<times> {..<M}"
       
  1310     let ?M = "Max (prod_encode ` ?P)"
       
  1311     have "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> (\<Sum>(m, n)\<in>?P. f (m, n))"
       
  1312       unfolding setsum_commute[of _ _ "{..<M}"] unfolding setsum_cartesian_product ..
       
  1313     also have "\<dots> \<le> (\<Sum>(m,n)\<in>(prod_decode ` {..?M}). f (m, n))"
       
  1314       by (auto intro!: setsum_mono3 image_eqI[where f=prod_decode, OF prod_encode_inverse[symmetric]])
       
  1315     also have "\<dots> \<le> y" using *[of "Suc ?M"]
       
  1316       by (simp add: lessThan_Suc_atMost[symmetric] setsum_reindex
       
  1317                inj_prod_decode del: setsum_lessThan_Suc)
       
  1318     finally show "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> y" .
       
  1319   qed
       
  1320 qed
       
  1321 
       
  1322 lemma pinfreal_mult_less_right:
       
  1323   assumes "b * a < c * a" "0 < a" "a < \<omega>"
       
  1324   shows "b < c"
       
  1325   using assms
       
  1326   by (cases a, cases b, cases c) (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
       
  1327 
       
  1328 lemma pinfreal_\<omega>_eq_plus[simp]: "\<omega> = a + b \<longleftrightarrow> (a = \<omega> \<or> b = \<omega>)"
       
  1329   by (cases a, cases b) auto
       
  1330 
       
  1331 lemma pinfreal_of_nat_le_iff:
       
  1332   "(of_nat k :: pinfreal) \<le> of_nat m \<longleftrightarrow> k \<le> m" by auto
       
  1333 
       
  1334 lemma pinfreal_of_nat_less_iff:
       
  1335   "(of_nat k :: pinfreal) < of_nat m \<longleftrightarrow> k < m" by auto
       
  1336 
       
  1337 lemma pinfreal_bound_add:
       
  1338   assumes "\<forall>N. f N + y \<le> (x::pinfreal)"
       
  1339   shows "(SUP n. f n) + y \<le> x"
       
  1340 proof (cases "x = \<omega>")
       
  1341   have "y \<le> x" using assms by (auto intro: pinfreal_le_add2)
       
  1342   assume "x \<noteq> \<omega>"
       
  1343   note move_y = pinfreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
       
  1344 
       
  1345   have "\<forall>N. f N \<le> x - y" using assms by (simp add: move_y)
       
  1346   hence "(SUP n. f n) \<le> x - y" by (rule SUPR_bound)
       
  1347   thus ?thesis by (simp add: move_y)
       
  1348 qed simp
       
  1349 
       
  1350 lemma SUPR_pinfreal_add:
       
  1351   fixes f g :: "nat \<Rightarrow> pinfreal"
       
  1352   assumes f: "\<forall>n. f n \<le> f (Suc n)" and g: "\<forall>n. g n \<le> g (Suc n)"
       
  1353   shows "(SUP n. f n + g n) = (SUP n. f n) + (SUP n. g n)"
       
  1354 proof (rule pinfreal_SUPI)
       
  1355   fix n :: nat from le_SUPI[of n UNIV f] le_SUPI[of n UNIV g]
       
  1356   show "f n + g n \<le> (SUP n. f n) + (SUP n. g n)"
       
  1357     by (auto intro!: add_mono)
       
  1358 next
       
  1359   fix y assume *: "\<And>n. n \<in> UNIV \<Longrightarrow> f n + g n \<le> y"
       
  1360   { fix n m
       
  1361     have "f n + g m \<le> y"
       
  1362     proof (cases rule: linorder_le_cases)
       
  1363       assume "n \<le> m"
       
  1364       hence "f n + g m \<le> f m + g m"
       
  1365         using f lift_Suc_mono_le by (auto intro!: add_right_mono)
       
  1366       also have "\<dots> \<le> y" using * by simp
       
  1367       finally show ?thesis .
       
  1368     next
       
  1369       assume "m \<le> n"
       
  1370       hence "f n + g m \<le> f n + g n"
       
  1371         using g lift_Suc_mono_le by (auto intro!: add_left_mono)
       
  1372       also have "\<dots> \<le> y" using * by simp
       
  1373       finally show ?thesis .
       
  1374     qed }
       
  1375   hence "\<And>m. \<forall>n. f n + g m \<le> y" by simp
       
  1376   from pinfreal_bound_add[OF this]
       
  1377   have "\<forall>m. (g m) + (SUP n. f n) \<le> y" by (simp add: ac_simps)
       
  1378   from pinfreal_bound_add[OF this]
       
  1379   show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
       
  1380 qed
       
  1381 
       
  1382 lemma SUPR_pinfreal_setsum:
       
  1383   fixes f :: "'x \<Rightarrow> nat \<Rightarrow> pinfreal"
       
  1384   assumes "\<And>i. i \<in> P \<Longrightarrow> \<forall>n. f i n \<le> f i (Suc n)"
       
  1385   shows "(SUP n. \<Sum>i\<in>P. f i n) = (\<Sum>i\<in>P. SUP n. f i n)"
       
  1386 proof cases
       
  1387   assume "finite P" from this assms show ?thesis
       
  1388   proof induct
       
  1389     case (insert i P)
       
  1390     thus ?case
       
  1391       apply simp
       
  1392       apply (subst SUPR_pinfreal_add)
       
  1393       by (auto intro!: setsum_mono)
       
  1394   qed simp
       
  1395 qed simp
       
  1396 
       
  1397 lemma Real_max:
       
  1398   assumes "x \<ge> 0" "y \<ge> 0"
       
  1399   shows "Real (max x y) = max (Real x) (Real y)"
       
  1400   using assms unfolding max_def by (auto simp add:not_le)
       
  1401 
       
  1402 lemma Real_real: "Real (real x) = (if x = \<omega> then 0 else x)"
       
  1403   using assms by (cases x) auto
       
  1404 
       
  1405 lemma inj_on_real: "inj_on real (UNIV - {\<omega>})"
       
  1406 proof (rule inj_onI)
       
  1407   fix x y assume mem: "x \<in> UNIV - {\<omega>}" "y \<in> UNIV - {\<omega>}" and "real x = real y"
       
  1408   thus "x = y" by (cases x, cases y) auto
       
  1409 qed
       
  1410 
       
  1411 lemma inj_on_Real: "inj_on Real {0..}"
       
  1412   by (auto intro!: inj_onI)
       
  1413 
       
  1414 lemma range_Real[simp]: "range Real = UNIV - {\<omega>}"
       
  1415 proof safe
       
  1416   fix x assume "x \<notin> range Real"
       
  1417   thus "x = \<omega>" by (cases x) auto
       
  1418 qed auto
       
  1419 
       
  1420 lemma image_Real[simp]: "Real ` {0..} = UNIV - {\<omega>}"
       
  1421 proof safe
       
  1422   fix x assume "x \<notin> Real ` {0..}"
       
  1423   thus "x = \<omega>" by (cases x) auto
       
  1424 qed auto
       
  1425 
       
  1426 lemma pinfreal_SUP_cmult:
       
  1427   fixes f :: "'a \<Rightarrow> pinfreal"
       
  1428   shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)"
       
  1429 proof (rule pinfreal_SUPI)
       
  1430   fix i assume "i \<in> R"
       
  1431   from le_SUPI[OF this]
       
  1432   show "z * f i \<le> z * (SUP i:R. f i)" by (rule pinfreal_mult_cancel)
       
  1433 next
       
  1434   fix y assume "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y"
       
  1435   hence *: "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y" by auto
       
  1436   show "z * (SUP i:R. f i) \<le> y"
       
  1437   proof (cases "\<forall>i\<in>R. f i = 0")
       
  1438     case True
       
  1439     show ?thesis
       
  1440     proof cases
       
  1441       assume "R \<noteq> {}" hence "f ` R = {0}" using True by auto
       
  1442       thus ?thesis by (simp add: SUPR_def)
       
  1443     qed (simp add: SUPR_def Sup_empty bot_pinfreal_def)
       
  1444   next
       
  1445     case False then obtain i where i: "i \<in> R" and f0: "f i \<noteq> 0" by auto
       
  1446     show ?thesis
       
  1447     proof (cases "z = 0 \<or> z = \<omega>")
       
  1448       case True with f0 *[OF i] show ?thesis by auto
       
  1449     next
       
  1450       case False hence z: "z \<noteq> 0" "z \<noteq> \<omega>" by auto
       
  1451       note div = pinfreal_inverse_le_eq[OF this, symmetric]
       
  1452       hence "\<And>i. i\<in>R \<Longrightarrow> f i \<le> y / z" using * by auto
       
  1453       thus ?thesis unfolding div SUP_le_iff by simp
       
  1454     qed
       
  1455   qed
       
  1456 qed
       
  1457 
       
  1458 instantiation pinfreal :: topological_space
       
  1459 begin
       
  1460 
       
  1461 definition "open A \<longleftrightarrow>
       
  1462   (\<exists>T. open T \<and> (Real ` (T\<inter>{0..}) = A - {\<omega>})) \<and> (\<omega> \<in> A \<longrightarrow> (\<exists>x\<ge>0. {Real x <..} \<subseteq> A))"
       
  1463 
       
  1464 lemma open_omega: "open A \<Longrightarrow> \<omega> \<in> A \<Longrightarrow> (\<exists>x\<ge>0. {Real x<..} \<subseteq> A)"
       
  1465   unfolding open_pinfreal_def by auto
       
  1466 
       
  1467 lemma open_omegaD: assumes "open A" "\<omega> \<in> A" obtains x where "x\<ge>0" "{Real x<..} \<subseteq> A"
       
  1468   using open_omega[OF assms] by auto
       
  1469 
       
  1470 lemma pinfreal_openE: assumes "open A" obtains A' x where
       
  1471   "open A'" "Real ` (A' \<inter> {0..}) = A - {\<omega>}"
       
  1472   "x \<ge> 0" "\<omega> \<in> A \<Longrightarrow> {Real x<..} \<subseteq> A"
       
  1473   using assms open_pinfreal_def by auto
       
  1474 
       
  1475 instance
       
  1476 proof
       
  1477   let ?U = "UNIV::pinfreal set"
       
  1478   show "open ?U" unfolding open_pinfreal_def
       
  1479     by (auto intro!: exI[of _ "UNIV"] exI[of _ 0])
       
  1480 next
       
  1481   fix S T::"pinfreal set" assume "open S" and "open T"
       
  1482   from `open S`[THEN pinfreal_openE] guess S' xS . note S' = this
       
  1483   from `open T`[THEN pinfreal_openE] guess T' xT . note T' = this
       
  1484 
       
  1485   from S'(1-3) T'(1-3)
       
  1486   show "open (S \<inter> T)" unfolding open_pinfreal_def
       
  1487   proof (safe intro!: exI[of _ "S' \<inter> T'"] exI[of _ "max xS xT"])
       
  1488     fix x assume *: "Real (max xS xT) < x" and "\<omega> \<in> S" "\<omega> \<in> T"
       
  1489     from `\<omega> \<in> S`[THEN S'(4)] * show "x \<in> S"
       
  1490       by (cases x, auto simp: max_def split: split_if_asm)
       
  1491     from `\<omega> \<in> T`[THEN T'(4)] * show "x \<in> T"
       
  1492       by (cases x, auto simp: max_def split: split_if_asm)
       
  1493   next
       
  1494     fix x assume x: "x \<notin> Real ` (S' \<inter> T' \<inter> {0..})"
       
  1495     have *: "S' \<inter> T' \<inter> {0..} = (S' \<inter> {0..}) \<inter> (T' \<inter> {0..})" by auto
       
  1496     assume "x \<in> T" "x \<in> S"
       
  1497     with S'(2) T'(2) show "x = \<omega>"
       
  1498       using x[unfolded *] inj_on_image_Int[OF inj_on_Real] by auto
       
  1499   qed auto
       
  1500 next
       
  1501   fix K assume openK: "\<forall>S \<in> K. open (S:: pinfreal set)"
       
  1502   hence "\<forall>S\<in>K. \<exists>T. open T \<and> Real ` (T \<inter> {0..}) = S - {\<omega>}" by (auto simp: open_pinfreal_def)
       
  1503   from bchoice[OF this] guess T .. note T = this[rule_format]
       
  1504 
       
  1505   show "open (\<Union>K)" unfolding open_pinfreal_def
       
  1506   proof (safe intro!: exI[of _ "\<Union>(T ` K)"])
       
  1507     fix x S assume "0 \<le> x" "x \<in> T S" "S \<in> K"
       
  1508     with T[OF `S \<in> K`] show "Real x \<in> \<Union>K" by auto
       
  1509   next
       
  1510     fix x S assume x: "x \<notin> Real ` (\<Union>T ` K \<inter> {0..})" "S \<in> K" "x \<in> S"
       
  1511     hence "x \<notin> Real ` (T S \<inter> {0..})"
       
  1512       by (auto simp: image_UN UN_simps[symmetric] simp del: UN_simps)
       
  1513     thus "x = \<omega>" using T[OF `S \<in> K`] `x \<in> S` by auto
       
  1514   next
       
  1515     fix S assume "\<omega> \<in> S" "S \<in> K"
       
  1516     from openK[rule_format, OF `S \<in> K`, THEN pinfreal_openE] guess S' x .
       
  1517     from this(3, 4) `\<omega> \<in> S`
       
  1518     show "\<exists>x\<ge>0. {Real x<..} \<subseteq> \<Union>K"
       
  1519       by (auto intro!: exI[of _ x] bexI[OF _ `S \<in> K`])
       
  1520   next
       
  1521     from T[THEN conjunct1] show "open (\<Union>T`K)" by auto
       
  1522   qed auto
       
  1523 qed
       
  1524 end
       
  1525 
       
  1526 lemma open_pinfreal_lessThan[simp]:
       
  1527   "open {..< a :: pinfreal}"
       
  1528 proof (cases a)
       
  1529   case (preal x) thus ?thesis unfolding open_pinfreal_def
       
  1530   proof (safe intro!: exI[of _ "{..< x}"])
       
  1531     fix y assume "y < Real x"
       
  1532     moreover assume "y \<notin> Real ` ({..<x} \<inter> {0..})"
       
  1533     ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
       
  1534     thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
       
  1535   qed auto
       
  1536 next
       
  1537   case infinite thus ?thesis
       
  1538     unfolding open_pinfreal_def by (auto intro!: exI[of _ UNIV])
       
  1539 qed
       
  1540 
       
  1541 lemma open_pinfreal_greaterThan[simp]:
       
  1542   "open {a :: pinfreal <..}"
       
  1543 proof (cases a)
       
  1544   case (preal x) thus ?thesis unfolding open_pinfreal_def
       
  1545   proof (safe intro!: exI[of _ "{x <..}"])
       
  1546     fix y assume "Real x < y"
       
  1547     moreover assume "y \<notin> Real ` ({x<..} \<inter> {0..})"
       
  1548     ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
       
  1549     thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
       
  1550   qed auto
       
  1551 next
       
  1552   case infinite thus ?thesis
       
  1553     unfolding open_pinfreal_def by (auto intro!: exI[of _ "{}"])
       
  1554 qed
       
  1555 
       
  1556 lemma pinfreal_open_greaterThanLessThan[simp]: "open {a::pinfreal <..< b}"
       
  1557   unfolding greaterThanLessThan_def by auto
       
  1558 
       
  1559 lemma closed_pinfreal_atLeast[simp, intro]: "closed {a :: pinfreal ..}"
       
  1560 proof -
       
  1561   have "- {a ..} = {..< a}" by auto
       
  1562   then show "closed {a ..}"
       
  1563     unfolding closed_def using open_pinfreal_lessThan by auto
       
  1564 qed
       
  1565 
       
  1566 lemma closed_pinfreal_atMost[simp, intro]: "closed {.. b :: pinfreal}"
       
  1567 proof -
       
  1568   have "- {.. b} = {b <..}" by auto
       
  1569   then show "closed {.. b}" 
       
  1570     unfolding closed_def using open_pinfreal_greaterThan by auto
       
  1571 qed
       
  1572 
       
  1573 lemma closed_pinfreal_atLeastAtMost[simp, intro]:
       
  1574   shows "closed {a :: pinfreal .. b}"
       
  1575   unfolding atLeastAtMost_def by auto
       
  1576 
       
  1577 lemma pinfreal_dense:
       
  1578   fixes x y :: pinfreal assumes "x < y"
       
  1579   shows "\<exists>z. x < z \<and> z < y"
       
  1580 proof -
       
  1581   from `x < y` obtain p where p: "x = Real p" "0 \<le> p" by (cases x) auto
       
  1582   show ?thesis
       
  1583   proof (cases y)
       
  1584     case (preal r) with p `x < y` have "p < r" by auto
       
  1585     with dense obtain z where "p < z" "z < r" by auto
       
  1586     thus ?thesis using preal p by (auto intro!: exI[of _ "Real z"])
       
  1587   next
       
  1588     case infinite thus ?thesis using `x < y` p
       
  1589       by (auto intro!: exI[of _ "Real p + 1"])
       
  1590   qed
       
  1591 qed
       
  1592 
       
  1593 instance pinfreal :: t2_space
       
  1594 proof
       
  1595   fix x y :: pinfreal assume "x \<noteq> y"
       
  1596   let "?P x (y::pinfreal)" = "\<exists> U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
       
  1597 
       
  1598   { fix x y :: pinfreal assume "x < y"
       
  1599     from pinfreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
       
  1600     have "?P x y"
       
  1601       apply (rule exI[of _ "{..<z}"])
       
  1602       apply (rule exI[of _ "{z<..}"])
       
  1603       using z by auto }
       
  1604   note * = this
       
  1605 
       
  1606   from `x \<noteq> y`
       
  1607   show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
       
  1608   proof (cases rule: linorder_cases)
       
  1609     assume "x = y" with `x \<noteq> y` show ?thesis by simp
       
  1610   next assume "x < y" from *[OF this] show ?thesis by auto
       
  1611   next assume "y < x" from *[OF this] show ?thesis by auto
       
  1612   qed
       
  1613 qed
       
  1614 
       
  1615 definition (in complete_lattice) isoton :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<up>" 50) where
       
  1616   "A \<up> X \<longleftrightarrow> (\<forall>i. A i \<le> A (Suc i)) \<and> (SUP i. A i) = X"
       
  1617 
       
  1618 definition (in complete_lattice) antiton (infix "\<down>" 50) where
       
  1619   "A \<down> X \<longleftrightarrow> (\<forall>i. A i \<ge> A (Suc i)) \<and> (INF i. A i) = X"
       
  1620 
       
  1621 lemma range_const[simp]: "range (\<lambda>x. c) = {c}" by auto
       
  1622 
       
  1623 lemma isoton_cmult_right:
       
  1624   assumes "f \<up> (x::pinfreal)"
       
  1625   shows "(\<lambda>i. c * f i) \<up> (c * x)"
       
  1626   using assms unfolding isoton_def pinfreal_SUP_cmult
       
  1627   by (auto intro: pinfreal_mult_cancel)
       
  1628 
       
  1629 lemma isoton_cmult_left:
       
  1630   "f \<up> (x::pinfreal) \<Longrightarrow> (\<lambda>i. f i * c) \<up> (x * c)"
       
  1631   by (subst (1 2) mult_commute) (rule isoton_cmult_right)
       
  1632 
       
  1633 lemma isoton_add:
       
  1634   assumes "f \<up> (x::pinfreal)" and "g \<up> y"
       
  1635   shows "(\<lambda>i. f i + g i) \<up> (x + y)"
       
  1636   using assms unfolding isoton_def
       
  1637   by (auto intro: pinfreal_mult_cancel add_mono simp: SUPR_pinfreal_add)
       
  1638 
       
  1639 lemma isoton_fun_expand:
       
  1640   "f \<up> x \<longleftrightarrow> (\<forall>i. (\<lambda>j. f j i) \<up> (x i))"
       
  1641 proof -
       
  1642   have "\<And>i. {y. \<exists>f'\<in>range f. y = f' i} = range (\<lambda>j. f j i)"
       
  1643     by auto
       
  1644   with assms show ?thesis
       
  1645     by (auto simp add: isoton_def le_fun_def Sup_fun_def SUPR_def)
       
  1646 qed
       
  1647 
       
  1648 lemma isoton_indicator:
       
  1649   assumes "f \<up> g"
       
  1650   shows "(\<lambda>i x. f i x * indicator A x) \<up> (\<lambda>x. g x * indicator A x :: pinfreal)"
       
  1651   using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
       
  1652 
       
  1653 lemma pinfreal_ord_one[simp]:
       
  1654   "Real p < 1 \<longleftrightarrow> p < 1"
       
  1655   "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
       
  1656   "1 < Real p \<longleftrightarrow> 1 < p"
       
  1657   "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
       
  1658   by (simp_all add: one_pinfreal_def del: Real_1)
       
  1659 
       
  1660 lemma SUP_mono:
       
  1661   assumes "\<And>n. f n \<le> g (N n)"
       
  1662   shows "(SUP n. f n) \<le> (SUP n. g n)"
       
  1663 proof (safe intro!: SUPR_bound)
       
  1664   fix n note assms[of n]
       
  1665   also have "g (N n) \<le> (SUP n. g n)" by (auto intro!: le_SUPI)
       
  1666   finally show "f n \<le> (SUP n. g n)" .
       
  1667 qed
       
  1668 
       
  1669 lemma isoton_Sup:
       
  1670   assumes "f \<up> u"
       
  1671   shows "f i \<le> u"
       
  1672   using le_SUPI[of i UNIV f] assms
       
  1673   unfolding isoton_def by auto
       
  1674 
       
  1675 lemma isoton_mono:
       
  1676   assumes iso: "x \<up> a" "y \<up> b" and *: "\<And>n. x n \<le> y (N n)"
       
  1677   shows "a \<le> b"
       
  1678 proof -
       
  1679   from iso have "a = (SUP n. x n)" "b = (SUP n. y n)"
       
  1680     unfolding isoton_def by auto
       
  1681   with * show ?thesis by (auto intro!: SUP_mono)
       
  1682 qed
       
  1683 
       
  1684 lemma pinfreal_le_mult_one_interval:
       
  1685   fixes x y :: pinfreal
       
  1686   assumes "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
       
  1687   shows "x \<le> y"
       
  1688 proof (cases x, cases y)
       
  1689   assume "x = \<omega>"
       
  1690   with assms[of "1 / 2"]
       
  1691   show "x \<le> y" by simp
       
  1692 next
       
  1693   fix r p assume *: "y = Real p" "x = Real r" and **: "0 \<le> r" "0 \<le> p"
       
  1694   have "r \<le> p"
       
  1695   proof (rule field_le_mult_one_interval)
       
  1696     fix z :: real assume "0 < z" and "z < 1"
       
  1697     with assms[of "Real z"]
       
  1698     show "z * r \<le> p" using ** * by (auto simp: zero_le_mult_iff)
       
  1699   qed
       
  1700   thus "x \<le> y" using ** * by simp
       
  1701 qed simp
       
  1702 
       
  1703 lemma pinfreal_greater_0[intro]:
       
  1704   fixes a :: pinfreal
       
  1705   assumes "a \<noteq> 0"
       
  1706   shows "a > 0"
       
  1707 using assms apply (cases a) by auto
       
  1708 
       
  1709 lemma pinfreal_mult_strict_right_mono:
       
  1710   assumes "a < b" and "0 < c" "c < \<omega>"
       
  1711   shows "a * c < b * c"
       
  1712   using assms
       
  1713   by (cases a, cases b, cases c)
       
  1714      (auto simp: zero_le_mult_iff pinfreal_less_\<omega>)
       
  1715 
       
  1716 lemma minus_pinfreal_eq2:
       
  1717   fixes x y z :: pinfreal
       
  1718   assumes "y \<le> x" and "y \<noteq> \<omega>" shows "z = x - y \<longleftrightarrow> z + y = x"
       
  1719   using assms
       
  1720   apply (subst eq_commute)
       
  1721   apply (subst minus_pinfreal_eq)
       
  1722   by (cases x, cases z, auto simp add: ac_simps not_less)
       
  1723 
       
  1724 lemma pinfreal_diff_eq_diff_imp_eq:
       
  1725   assumes "a \<noteq> \<omega>" "b \<le> a" "c \<le> a"
       
  1726   assumes "a - b = a - c"
       
  1727   shows "b = c"
       
  1728   using assms
       
  1729   by (cases a, cases b, cases c) (auto split: split_if_asm)
       
  1730 
       
  1731 lemma pinfreal_inverse_eq_0: "inverse x = 0 \<longleftrightarrow> x = \<omega>"
       
  1732   by (cases x) auto
       
  1733 
       
  1734 lemma pinfreal_mult_inverse:
       
  1735   "\<lbrakk> x \<noteq> \<omega> ; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x * inverse x = 1"
       
  1736   by (cases x) auto
       
  1737 
       
  1738 lemma pinfreal_zero_less_diff_iff:
       
  1739   fixes a b :: pinfreal shows "0 < a - b \<longleftrightarrow> b < a"
       
  1740   apply (cases a, cases b)
       
  1741   apply (auto simp: pinfreal_noteq_omega_Ex pinfreal_less_\<omega>)
       
  1742   apply (cases b)
       
  1743   by auto
       
  1744 
       
  1745 lemma pinfreal_less_Real_Ex:
       
  1746   fixes a b :: pinfreal shows "x < Real r \<longleftrightarrow> (\<exists>p\<ge>0. p < r \<and> x = Real p)"
       
  1747   by (cases x) auto
       
  1748 
       
  1749 lemma open_Real: assumes "open S" shows "open (Real ` ({0..} \<inter> S))"
       
  1750   unfolding open_pinfreal_def apply(rule,rule,rule,rule assms) by auto
       
  1751 
       
  1752 lemma pinfreal_zero_le_diff:
       
  1753   fixes a b :: pinfreal shows "a - b = 0 \<longleftrightarrow> a \<le> b"
       
  1754   by (cases a, cases b, simp_all, cases b, auto)
       
  1755 
       
  1756 lemma lim_Real[simp]: assumes "\<forall>n. f n \<ge> 0" "m\<ge>0"
       
  1757   shows "(\<lambda>n. Real (f n)) ----> Real m \<longleftrightarrow> (\<lambda>n. f n) ----> m" (is "?l = ?r")
       
  1758 proof assume ?l show ?r unfolding Lim_sequentially
       
  1759   proof safe fix e::real assume e:"e>0"
       
  1760     note open_ball[of m e] note open_Real[OF this]
       
  1761     note * = `?l`[unfolded tendsto_def,rule_format,OF this]
       
  1762     have "eventually (\<lambda>x. Real (f x) \<in> Real ` ({0..} \<inter> ball m e)) sequentially"
       
  1763       apply(rule *) unfolding image_iff using assms(2) e by auto
       
  1764     thus "\<exists>N. \<forall>n\<ge>N. dist (f n) m < e" unfolding eventually_sequentially 
       
  1765       apply safe apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
       
  1766     proof- fix n x assume "Real (f n) = Real x" "0 \<le> x"
       
  1767       hence *:"f n = x" using assms(1) by auto
       
  1768       assume "x \<in> ball m e" thus "dist (f n) m < e" unfolding *
       
  1769         by (auto simp add:dist_commute)
       
  1770     qed qed
       
  1771 next assume ?r show ?l unfolding tendsto_def eventually_sequentially 
       
  1772   proof safe fix S assume S:"open S" "Real m \<in> S"
       
  1773     guess T y using S(1) apply-apply(erule pinfreal_openE) . note T=this
       
  1774     have "m\<in>real ` (S - {\<omega>})" unfolding image_iff 
       
  1775       apply(rule_tac x="Real m" in bexI) using assms(2) S(2) by auto
       
  1776     hence "m \<in> T" unfolding T(2)[THEN sym] by auto 
       
  1777     from `?r`[unfolded tendsto_def eventually_sequentially,rule_format,OF T(1) this]
       
  1778     guess N .. note N=this[rule_format]
       
  1779     show "\<exists>N. \<forall>n\<ge>N. Real (f n) \<in> S" apply(rule_tac x=N in exI) 
       
  1780     proof safe fix n assume n:"N\<le>n"
       
  1781       have "f n \<in> real ` (S - {\<omega>})" using N[OF n] assms unfolding T(2)[THEN sym] 
       
  1782         unfolding image_iff apply-apply(rule_tac x="Real (f n)" in bexI)
       
  1783         unfolding real_Real by auto
       
  1784       then guess x unfolding image_iff .. note x=this
       
  1785       show "Real (f n) \<in> S" unfolding x apply(subst Real_real) using x by auto
       
  1786     qed
       
  1787   qed
       
  1788 qed
       
  1789 
       
  1790 lemma pinfreal_INFI:
       
  1791   fixes x :: pinfreal
       
  1792   assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
       
  1793   assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x"
       
  1794   shows "(INF i:A. f i) = x"
       
  1795   unfolding INFI_def Inf_pinfreal_def
       
  1796   using assms by (auto intro!: Greatest_equality)
       
  1797 
       
  1798 lemma real_of_pinfreal_less:"x < y \<Longrightarrow> y\<noteq>\<omega> \<Longrightarrow> real x < real y"
       
  1799 proof- case goal1
       
  1800   have *:"y = Real (real y)" "x = Real (real x)" using goal1 Real_real by auto
       
  1801   show ?case using goal1 apply- apply(subst(asm) *(1))apply(subst(asm) *(2))
       
  1802     unfolding pinfreal_less by auto
       
  1803 qed
       
  1804 
       
  1805 lemma not_less_omega[simp]:"\<not> x < \<omega> \<longleftrightarrow> x = \<omega>"
       
  1806   by (metis antisym_conv3 pinfreal_less(3)) 
       
  1807 
       
  1808 lemma Real_real': assumes "x\<noteq>\<omega>" shows "Real (real x) = x"
       
  1809 proof- have *:"(THE r. 0 \<le> r \<and> x = Real r) = real x"
       
  1810     apply(rule the_equality) using assms unfolding Real_real by auto
       
  1811   have "Real (THE r. 0 \<le> r \<and> x = Real r) = x" unfolding *
       
  1812     using assms unfolding Real_real by auto
       
  1813   thus ?thesis unfolding real_of_pinfreal_def of_pinfreal_def
       
  1814     unfolding pinfreal_case_def using assms by auto
       
  1815 qed 
       
  1816 
       
  1817 lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)" 
       
  1818   unfolding pinfreal_less by auto
       
  1819 
       
  1820 lemma Lim_omega: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
       
  1821 proof assume ?r show ?l apply(rule topological_tendstoI)
       
  1822     unfolding eventually_sequentially
       
  1823   proof- fix S assume "open S" "\<omega> \<in> S"
       
  1824     from open_omega[OF this] guess B .. note B=this
       
  1825     from `?r`[rule_format,of "(max B 0)+1"] guess N .. note N=this
       
  1826     show "\<exists>N. \<forall>n\<ge>N. f n \<in> S" apply(rule_tac x=N in exI)
       
  1827     proof safe case goal1 
       
  1828       have "Real B < Real ((max B 0) + 1)" by auto
       
  1829       also have "... \<le> f n" using goal1 N by auto
       
  1830       finally show ?case using B by fastsimp
       
  1831     qed
       
  1832   qed
       
  1833 next assume ?l show ?r
       
  1834   proof fix B::real have "open {Real B<..}" "\<omega> \<in> {Real B<..}" by auto
       
  1835     from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
       
  1836     guess N .. note N=this
       
  1837     show "\<exists>N. \<forall>n\<ge>N. Real B \<le> f n" apply(rule_tac x=N in exI) using N by auto
       
  1838   qed
       
  1839 qed
       
  1840 
       
  1841 lemma Lim_bounded_omgea: assumes lim:"f ----> l" and "\<And>n. f n \<le> Real B" shows "l \<noteq> \<omega>"
       
  1842 proof(rule ccontr,unfold not_not) let ?B = "max (B + 1) 1" assume as:"l=\<omega>"
       
  1843   from lim[unfolded this Lim_omega,rule_format,of "?B"]
       
  1844   guess N .. note N=this[rule_format,OF le_refl]
       
  1845   hence "Real ?B \<le> Real B" using assms(2)[of N] by(rule order_trans) 
       
  1846   hence "Real ?B < Real ?B" using Real_less_plus_one[of B] by(rule le_less_trans)
       
  1847   thus False by auto
       
  1848 qed
       
  1849 
       
  1850 lemma incseq_le_pinfreal: assumes inc: "\<And>n m. n\<ge>m \<Longrightarrow> X n \<ge> X m"
       
  1851   and lim: "X ----> (L::pinfreal)" shows "X n \<le> L"
       
  1852 proof(cases "L = \<omega>")
       
  1853   case False have "\<forall>n. X n \<noteq> \<omega>"
       
  1854   proof(rule ccontr,unfold not_all not_not,safe)
       
  1855     case goal1 hence "\<forall>n\<ge>x. X n = \<omega>" using inc[of x] by auto
       
  1856     hence "X ----> \<omega>" unfolding tendsto_def eventually_sequentially
       
  1857       apply safe apply(rule_tac x=x in exI) by auto
       
  1858     note Lim_unique[OF trivial_limit_sequentially this lim]
       
  1859     with False show False by auto
       
  1860   qed note * =this[rule_format]
       
  1861 
       
  1862   have **:"\<forall>m n. m \<le> n \<longrightarrow> Real (real (X m)) \<le> Real (real (X n))"
       
  1863     unfolding Real_real using * inc by auto
       
  1864   have "real (X n) \<le> real L" apply-apply(rule incseq_le) defer
       
  1865     apply(subst lim_Real[THEN sym]) apply(rule,rule,rule)
       
  1866     unfolding Real_real'[OF *] Real_real'[OF False] 
       
  1867     unfolding incseq_def using ** lim by auto
       
  1868   hence "Real (real (X n)) \<le> Real (real L)" by auto
       
  1869   thus ?thesis unfolding Real_real using * False by auto
       
  1870 qed auto
       
  1871 
       
  1872 lemma SUP_Lim_pinfreal: assumes "\<And>n m. n\<ge>m \<Longrightarrow> f n \<ge> f m" "f ----> l"
       
  1873   shows "(SUP n. f n) = (l::pinfreal)" unfolding SUPR_def Sup_pinfreal_def
       
  1874 proof (safe intro!: Least_equality)
       
  1875   fix n::nat show "f n \<le> l" apply(rule incseq_le_pinfreal)
       
  1876     using assms by auto
       
  1877 next fix y assume y:"\<forall>x\<in>range f. x \<le> y" show "l \<le> y"
       
  1878   proof(rule ccontr,cases "y=\<omega>",unfold not_le)
       
  1879     case False assume as:"y < l"
       
  1880     have l:"l \<noteq> \<omega>" apply(rule Lim_bounded_omgea[OF assms(2), of "real y"])
       
  1881       using False y unfolding Real_real by auto
       
  1882 
       
  1883     have yl:"real y < real l" using as apply-
       
  1884       apply(subst(asm) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
       
  1885       apply(subst(asm) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) 
       
  1886       unfolding pinfreal_less apply(subst(asm) if_P) by auto
       
  1887     hence "y + (y - l) * Real (1 / 2) < l" apply-
       
  1888       apply(subst Real_real'[THEN sym,OF `y\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
       
  1889       apply(subst Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) by auto
       
  1890     hence *:"l \<in> {y + (y - l) / 2<..}" by auto
       
  1891     have "open {y + (y-l)/2 <..}" by auto
       
  1892     note topological_tendstoD[OF assms(2) this *]
       
  1893     from this[unfolded eventually_sequentially] guess N .. note this[rule_format, of N]
       
  1894     hence "y + (y - l) * Real (1 / 2) < y" using y[rule_format,of "f N"] by auto
       
  1895     hence "Real (real y) + (Real (real y) - Real (real l)) * Real (1 / 2) < Real (real y)"
       
  1896       unfolding Real_real using `y\<noteq>\<omega>` `l\<noteq>\<omega>` by auto
       
  1897     thus False using yl by auto
       
  1898   qed auto
       
  1899 qed
       
  1900 
       
  1901 lemma Real_max':"Real x = Real (max x 0)" 
       
  1902 proof(cases "x < 0") case True
       
  1903   hence *:"max x 0 = 0" by auto
       
  1904   show ?thesis unfolding * using True by auto
       
  1905 qed auto
       
  1906 
       
  1907 lemma lim_pinfreal_increasing: assumes "\<forall>n m. n\<ge>m \<longrightarrow> f n \<ge> f m"
       
  1908   obtains l where "f ----> (l::pinfreal)"
       
  1909 proof(cases "\<exists>B. \<forall>n. f n < Real B")
       
  1910   case False thus thesis apply- apply(rule that[of \<omega>]) unfolding Lim_omega not_ex not_all
       
  1911     apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
       
  1912     apply(rule order_trans[OF _ assms[rule_format]]) by auto
       
  1913 next case True then guess B .. note B = this[rule_format]
       
  1914   hence *:"\<And>n. f n < \<omega>" apply-apply(rule less_le_trans,assumption) by auto
       
  1915   have *:"\<And>n. f n \<noteq> \<omega>" proof- case goal1 show ?case using *[of n] by auto qed
       
  1916   have B':"\<And>n. real (f n) \<le> max 0 B" proof- case goal1 thus ?case
       
  1917       using B[of n] apply-apply(subst(asm) Real_real'[THEN sym]) defer
       
  1918       apply(subst(asm)(2) Real_max') unfolding pinfreal_less apply(subst(asm) if_P) using *[of n] by auto
       
  1919   qed
       
  1920   have "\<exists>l. (\<lambda>n. real (f n)) ----> l" apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
       
  1921   proof safe show "bounded {real (f n) |n. True}"
       
  1922       unfolding bounded_def apply(rule_tac x=0 in exI,rule_tac x="max 0 B" in exI)
       
  1923       using B' unfolding dist_norm by auto
       
  1924     fix n::nat have "Real (real (f n)) \<le> Real (real (f (Suc n)))"
       
  1925       using assms[rule_format,of n "Suc n"] apply(subst Real_real)+
       
  1926       using *[of n] *[of "Suc n"] by fastsimp
       
  1927     thus "real (f n) \<le> real (f (Suc n))" by auto
       
  1928   qed then guess l .. note l=this
       
  1929   have "0 \<le> l" apply(rule LIMSEQ_le_const[OF l])
       
  1930     by(rule_tac x=0 in exI,auto)
       
  1931 
       
  1932   thus ?thesis apply-apply(rule that[of "Real l"])
       
  1933     using l apply-apply(subst(asm) lim_Real[THEN sym]) prefer 3
       
  1934     unfolding Real_real using * by auto
       
  1935 qed
       
  1936 
       
  1937 lemma setsum_neq_omega: assumes "finite s" "\<And>x. x \<in> s \<Longrightarrow> f x \<noteq> \<omega>"
       
  1938   shows "setsum f s \<noteq> \<omega>" using assms
       
  1939 proof induct case (insert x s)
       
  1940   show ?case unfolding setsum.insert[OF insert(1-2)] 
       
  1941     using insert by auto
       
  1942 qed auto
       
  1943 
       
  1944 
       
  1945 lemma real_Real': "0 \<le> x \<Longrightarrow> real (Real x) = x"
       
  1946   unfolding real_Real by auto
       
  1947 
       
  1948 lemma real_pinfreal_pos[intro]:
       
  1949   assumes "x \<noteq> 0" "x \<noteq> \<omega>"
       
  1950   shows "real x > 0"
       
  1951   apply(subst real_Real'[THEN sym,of 0]) defer
       
  1952   apply(rule real_of_pinfreal_less) using assms by auto
       
  1953 
       
  1954 lemma Lim_omega_gt: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n > Real B)" (is "?l = ?r")
       
  1955 proof assume ?l thus ?r unfolding Lim_omega apply safe
       
  1956     apply(erule_tac x="max B 0 +1" in allE,safe)
       
  1957     apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
       
  1958     apply(rule_tac y="Real (max B 0 + 1)" in less_le_trans) by auto
       
  1959 next assume ?r thus ?l unfolding Lim_omega apply safe
       
  1960     apply(erule_tac x=B in allE,safe) apply(rule_tac x=N in exI,safe) by auto
       
  1961 qed
       
  1962 
       
  1963 lemma pinfreal_minus_le_cancel:
       
  1964   fixes a b c :: pinfreal
       
  1965   assumes "b \<le> a"
       
  1966   shows "c - a \<le> c - b"
       
  1967   using assms by (cases a, cases b, cases c, simp, simp, simp, cases b, cases c, simp_all)
       
  1968 
       
  1969 lemma pinfreal_minus_\<omega>[simp]: "x - \<omega> = 0" by (cases x) simp_all
       
  1970 
       
  1971 lemma pinfreal_minus_mono[intro]: "a - x \<le> (a::pinfreal)"
       
  1972 proof- have "a - x \<le> a - 0"
       
  1973     apply(rule pinfreal_minus_le_cancel) by auto
       
  1974   thus ?thesis by auto
       
  1975 qed
       
  1976 
       
  1977 lemma pinfreal_minus_eq_\<omega>[simp]: "x - y = \<omega> \<longleftrightarrow> (x = \<omega> \<and> y \<noteq> \<omega>)"
       
  1978   by (cases x, cases y) (auto, cases y, auto)
       
  1979 
       
  1980 lemma pinfreal_less_minus_iff:
       
  1981   fixes a b c :: pinfreal
       
  1982   shows "a < b - c \<longleftrightarrow> c + a < b"
       
  1983   by (cases c, cases a, cases b, auto)
       
  1984 
       
  1985 lemma pinfreal_minus_less_iff:
       
  1986   fixes a b c :: pinfreal shows "a - c < b \<longleftrightarrow> (0 < b \<and> (c \<noteq> \<omega> \<longrightarrow> a < b + c))"
       
  1987   by (cases c, cases a, cases b, auto)
       
  1988 
       
  1989 lemma pinfreal_le_minus_iff:
       
  1990   fixes a b c :: pinfreal
       
  1991   shows "a \<le> c - b \<longleftrightarrow> ((c \<le> b \<longrightarrow> a = 0) \<and> (b < c \<longrightarrow> a + b \<le> c))"
       
  1992   by (cases a, cases c, cases b, auto simp: pinfreal_noteq_omega_Ex)
       
  1993 
       
  1994 lemma pinfreal_minus_le_iff:
       
  1995   fixes a b c :: pinfreal
       
  1996   shows "a - c \<le> b \<longleftrightarrow> (c \<le> a \<longrightarrow> a \<le> b + c)"
       
  1997   by (cases a, cases c, cases b, auto simp: pinfreal_noteq_omega_Ex)
       
  1998 
       
  1999 lemmas pinfreal_minus_order = pinfreal_minus_le_iff pinfreal_minus_less_iff pinfreal_le_minus_iff pinfreal_less_minus_iff
       
  2000 
       
  2001 lemma pinfreal_minus_strict_mono:
       
  2002   assumes "a > 0" "x > 0" "a\<noteq>\<omega>"
       
  2003   shows "a - x < (a::pinfreal)"
       
  2004   using assms by(cases x, cases a, auto)
       
  2005 
       
  2006 lemma pinfreal_minus':
       
  2007   "Real r - Real p = (if 0 \<le> r \<and> p \<le> r then if 0 \<le> p then Real (r - p) else Real r else 0)"
       
  2008   by (auto simp: minus_pinfreal_eq not_less)
       
  2009 
       
  2010 lemma pinfreal_minus_plus:
       
  2011   "x \<le> (a::pinfreal) \<Longrightarrow> a - x + x = a"
       
  2012   by (cases a, cases x) auto
       
  2013 
       
  2014 lemma pinfreal_cancel_plus_minus: "b \<noteq> \<omega> \<Longrightarrow> a + b - b = a"
       
  2015   by (cases a, cases b) auto
       
  2016 
       
  2017 lemma pinfreal_minus_le_cancel_right:
       
  2018   fixes a b c :: pinfreal
       
  2019   assumes "a \<le> b" "c \<le> a"
       
  2020   shows "a - c \<le> b - c"
       
  2021   using assms by (cases a, cases b, cases c, auto, cases c, auto)
       
  2022 
       
  2023 lemma real_of_pinfreal_setsum':
       
  2024   assumes "\<forall>x \<in> S. f x \<noteq> \<omega>"
       
  2025   shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
       
  2026 proof cases
       
  2027   assume "finite S"
       
  2028   from this assms show ?thesis
       
  2029     by induct (simp_all add: real_of_pinfreal_add setsum_\<omega>)
       
  2030 qed simp
       
  2031 
       
  2032 lemma Lim_omega_pos: "f ----> \<omega> \<longleftrightarrow> (\<forall>B>0. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
       
  2033   unfolding Lim_omega apply safe defer
       
  2034   apply(erule_tac x="max 1 B" in allE) apply safe defer
       
  2035   apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
       
  2036   apply(rule_tac y="Real (max 1 B)" in order_trans) by auto
       
  2037 
       
  2038 lemma (in complete_lattice) isotonD[dest]:
       
  2039   assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
       
  2040   using assms unfolding isoton_def by auto
       
  2041 
       
  2042 lemma isotonD'[dest]:
       
  2043   assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
       
  2044   using assms unfolding isoton_def le_fun_def by auto
       
  2045 
       
  2046 lemma pinfreal_LimI_finite:
       
  2047   assumes "x \<noteq> \<omega>" "\<And>r. 0 < r \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
       
  2048   shows "u ----> x"
       
  2049 proof (rule topological_tendstoI, unfold eventually_sequentially)
       
  2050   fix S assume "open S" "x \<in> S"
       
  2051   then obtain A where "open A" and A_eq: "Real ` (A \<inter> {0..}) = S - {\<omega>}" by (auto elim!: pinfreal_openE)
       
  2052   then have "x \<in> Real ` (A \<inter> {0..})" using `x \<in> S` `x \<noteq> \<omega>` by auto
       
  2053   then have "real x \<in> A" by auto
       
  2054   then obtain r where "0 < r" and dist: "\<And>y. dist y (real x) < r \<Longrightarrow> y \<in> A"
       
  2055     using `open A` unfolding open_real_def by auto
       
  2056   then obtain n where
       
  2057     upper: "\<And>N. n \<le> N \<Longrightarrow> u N < x + Real r" and
       
  2058     lower: "\<And>N. n \<le> N \<Longrightarrow> x < u N + Real r" using assms(2)[of "Real r"] by auto
       
  2059   show "\<exists>N. \<forall>n\<ge>N. u n \<in> S"
       
  2060   proof (safe intro!: exI[of _ n])
       
  2061     fix N assume "n \<le> N"
       
  2062     from upper[OF this] `x \<noteq> \<omega>` `0 < r`
       
  2063     have "u N \<noteq> \<omega>" by (force simp: pinfreal_noteq_omega_Ex)
       
  2064     with `x \<noteq> \<omega>` `0 < r` lower[OF `n \<le> N`] upper[OF `n \<le> N`]
       
  2065     have "dist (real (u N)) (real x) < r" "u N \<noteq> \<omega>"
       
  2066       by (auto simp: pinfreal_noteq_omega_Ex dist_real_def abs_diff_less_iff field_simps)
       
  2067     from dist[OF this(1)]
       
  2068     have "u N \<in> Real ` (A \<inter> {0..})" using `u N \<noteq> \<omega>`
       
  2069       by (auto intro!: image_eqI[of _ _ "real (u N)"] simp: pinfreal_noteq_omega_Ex Real_real)
       
  2070     thus "u N \<in> S" using A_eq by simp
       
  2071   qed
       
  2072 qed
       
  2073 
       
  2074 lemma real_Real_max:"real (Real x) = max x 0"
       
  2075   unfolding real_Real by auto
       
  2076 
       
  2077 lemma (in complete_lattice) SUPR_upper:
       
  2078   "x \<in> A \<Longrightarrow> f x \<le> SUPR A f"
       
  2079   unfolding SUPR_def apply(rule Sup_upper) by auto
       
  2080 
       
  2081 lemma (in complete_lattice) SUPR_subset:
       
  2082   assumes "A \<subseteq> B" shows "SUPR A f \<le> SUPR B f"
       
  2083   apply(rule SUP_leI) apply(rule SUPR_upper) using assms by auto
       
  2084 
       
  2085 lemma (in complete_lattice) SUPR_mono:
       
  2086   assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
       
  2087   shows "SUPR A f \<le> SUPR B f"
       
  2088   unfolding SUPR_def apply(rule Sup_mono)
       
  2089   using assms by auto
       
  2090 
       
  2091 lemma Sup_lim:
       
  2092   assumes "\<forall>n. b n \<in> s" "b ----> (a::pinfreal)"
       
  2093   shows "a \<le> Sup s"
       
  2094 proof(rule ccontr,unfold not_le)
       
  2095   assume as:"Sup s < a" hence om:"Sup s \<noteq> \<omega>" by auto
       
  2096   have s:"s \<noteq> {}" using assms by auto
       
  2097   { presume *:"\<forall>n. b n < a \<Longrightarrow> False"
       
  2098     show False apply(cases,rule *,assumption,unfold not_all not_less)
       
  2099     proof- case goal1 then guess n .. note n=this
       
  2100       thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]]
       
  2101         using as by auto
       
  2102     qed
       
  2103   } assume b:"\<forall>n. b n < a"
       
  2104   show False
       
  2105   proof(cases "a = \<omega>")
       
  2106     case False have *:"a - Sup s > 0" 
       
  2107       using False as by(auto simp: pinfreal_zero_le_diff)
       
  2108     have "(a - Sup s) / 2 \<le> a / 2" unfolding divide_pinfreal_def
       
  2109       apply(rule mult_right_mono) by auto
       
  2110     also have "... = Real (real (a / 2))" apply(rule Real_real'[THEN sym])
       
  2111       using False by auto
       
  2112     also have "... < Real (real a)" unfolding pinfreal_less using as False
       
  2113       by(auto simp add: real_of_pinfreal_mult[THEN sym])
       
  2114     also have "... = a" apply(rule Real_real') using False by auto
       
  2115     finally have asup:"a > (a - Sup s) / 2" .
       
  2116     have "\<exists>n. a - b n < (a - Sup s) / 2"
       
  2117     proof(rule ccontr,unfold not_ex not_less)
       
  2118       case goal1
       
  2119       have "(a - Sup s) * Real (1 / 2)  > 0" 
       
  2120         using * by auto
       
  2121       hence "a - (a - Sup s) * Real (1 / 2) < a"
       
  2122         apply-apply(rule pinfreal_minus_strict_mono)
       
  2123         using False * by auto
       
  2124       hence *:"a \<in> {a - (a - Sup s) / 2<..}"using asup by auto 
       
  2125       note topological_tendstoD[OF assms(2) open_pinfreal_greaterThan,OF *]
       
  2126       from this[unfolded eventually_sequentially] guess n .. 
       
  2127       note n = this[rule_format,of n] 
       
  2128       have "b n + (a - Sup s) / 2 \<le> a" 
       
  2129         using add_right_mono[OF goal1[rule_format,of n],of "b n"]
       
  2130         unfolding pinfreal_minus_plus[OF less_imp_le[OF b[rule_format]]]
       
  2131         by(auto simp: add_commute)
       
  2132       hence "b n \<le> a - (a - Sup s) / 2" unfolding pinfreal_le_minus_iff
       
  2133         using asup by auto
       
  2134       hence "b n \<notin> {a - (a - Sup s) / 2<..}" by auto
       
  2135       thus False using n by auto
       
  2136     qed
       
  2137     then guess n .. note n = this
       
  2138     have "Sup s < a - (a - Sup s) / 2"
       
  2139       using False as om by (cases a) (auto simp: pinfreal_noteq_omega_Ex field_simps)
       
  2140     also have "... \<le> b n"
       
  2141     proof- note add_right_mono[OF less_imp_le[OF n],of "b n"]
       
  2142       note this[unfolded pinfreal_minus_plus[OF less_imp_le[OF b[rule_format]]]]
       
  2143       hence "a - (a - Sup s) / 2 \<le> (a - Sup s) / 2 + b n - (a - Sup s) / 2"
       
  2144         apply(rule pinfreal_minus_le_cancel_right) using asup by auto
       
  2145       also have "... = b n + (a - Sup s) / 2 - (a - Sup s) / 2" 
       
  2146         by(auto simp add: add_commute)
       
  2147       also have "... = b n" apply(subst pinfreal_cancel_plus_minus)
       
  2148       proof(rule ccontr,unfold not_not) case goal1
       
  2149         show ?case using asup unfolding goal1 by auto 
       
  2150       qed auto
       
  2151       finally show ?thesis .
       
  2152     qed     
       
  2153     finally show False
       
  2154       using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] by auto  
       
  2155   next case True
       
  2156     from assms(2)[unfolded True Lim_omega_gt,rule_format,of "real (Sup s)"]
       
  2157     guess N .. note N = this[rule_format,of N]
       
  2158     thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of N]] 
       
  2159       unfolding Real_real using om by auto
       
  2160   qed qed
       
  2161 
       
  2162 lemma less_SUP_iff:
       
  2163   fixes a :: pinfreal
       
  2164   shows "(a < (SUP i:A. f i)) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
       
  2165   unfolding SUPR_def less_Sup_iff by auto
       
  2166 
       
  2167 lemma Sup_mono_lim:
       
  2168   assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> b ----> (a::pinfreal)"
       
  2169   shows "Sup A \<le> Sup B"
       
  2170   unfolding Sup_le_iff apply(rule) apply(drule assms[rule_format]) apply safe
       
  2171   apply(rule_tac b=b in Sup_lim) by auto
       
  2172 
       
  2173 lemma pinfreal_less_add:
       
  2174   assumes "x \<noteq> \<omega>" "a < b"
       
  2175   shows "x + a < x + b"
       
  2176   using assms by (cases a, cases b, cases x) auto
       
  2177 
       
  2178 lemma SUPR_lim:
       
  2179   assumes "\<forall>n. b n \<in> B" "(\<lambda>n. f (b n)) ----> (f a::pinfreal)"
       
  2180   shows "f a \<le> SUPR B f"
       
  2181   unfolding SUPR_def apply(rule Sup_lim[of "\<lambda>n. f (b n)"])
       
  2182   using assms by auto
       
  2183 
       
  2184 lemma SUP_\<omega>_imp:
       
  2185   assumes "(SUP i. f i) = \<omega>"
       
  2186   shows "\<exists>i. Real x < f i"
       
  2187 proof (rule ccontr)
       
  2188   assume "\<not> ?thesis" hence "\<And>i. f i \<le> Real x" by (simp add: not_less)
       
  2189   hence "(SUP i. f i) \<le> Real x" unfolding SUP_le_iff by auto
       
  2190   with assms show False by auto
       
  2191 qed
       
  2192 
       
  2193 lemma SUPR_mono_lim:
       
  2194   assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> (\<lambda>n. f (b n)) ----> (f a::pinfreal)"
       
  2195   shows "SUPR A f \<le> SUPR B f"
       
  2196   unfolding SUPR_def apply(rule Sup_mono_lim)
       
  2197   apply safe apply(drule assms[rule_format],safe)
       
  2198   apply(rule_tac x="\<lambda>n. f (b n)" in exI) by auto
       
  2199 
       
  2200 lemma real_0_imp_eq_0:
       
  2201   assumes "x \<noteq> \<omega>" "real x = 0"
       
  2202   shows "x = 0"
       
  2203 using assms by (cases x) auto
       
  2204 
       
  2205 lemma SUPR_mono:
       
  2206   assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
       
  2207   shows "SUPR A f \<le> SUPR B f"
       
  2208   unfolding SUPR_def apply(rule Sup_mono)
       
  2209   using assms by auto
       
  2210 
       
  2211 lemma less_add_Real:
       
  2212   fixes x :: real
       
  2213   fixes a b :: pinfreal
       
  2214   assumes "x \<ge> 0" "a < b"
       
  2215   shows "a + Real x < b + Real x"
       
  2216 using assms by (cases a, cases b) auto
       
  2217 
       
  2218 lemma le_add_Real:
       
  2219   fixes x :: real
       
  2220   fixes a b :: pinfreal
       
  2221   assumes "x \<ge> 0" "a \<le> b"
       
  2222   shows "a + Real x \<le> b + Real x"
       
  2223 using assms by (cases a, cases b) auto
       
  2224 
       
  2225 lemma le_imp_less_pinfreal:
       
  2226   fixes x :: pinfreal
       
  2227   assumes "x > 0" "a + x \<le> b" "a \<noteq> \<omega>"
       
  2228   shows "a < b"
       
  2229 using assms by (cases x, cases a, cases b) auto
       
  2230 
       
  2231 lemma pinfreal_INF_minus:
       
  2232   fixes f :: "nat \<Rightarrow> pinfreal"
       
  2233   assumes "c \<noteq> \<omega>"
       
  2234   shows "(INF i. c - f i) = c - (SUP i. f i)"
       
  2235 proof (cases "SUP i. f i")
       
  2236   case infinite
       
  2237   from `c \<noteq> \<omega>` obtain x where [simp]: "c = Real x" by (cases c) auto
       
  2238   from SUP_\<omega>_imp[OF infinite] obtain i where "Real x < f i" by auto
       
  2239   have "(INF i. c - f i) \<le> c - f i"
       
  2240     by (auto intro!: complete_lattice_class.INF_leI)
       
  2241   also have "\<dots> = 0" using `Real x < f i` by (auto simp: minus_pinfreal_eq)
       
  2242   finally show ?thesis using infinite by auto
       
  2243 next
       
  2244   case (preal r)
       
  2245   from `c \<noteq> \<omega>` obtain x where c: "c = Real x" by (cases c) auto
       
  2246 
       
  2247   show ?thesis unfolding c
       
  2248   proof (rule pinfreal_INFI)
       
  2249     fix i have "f i \<le> (SUP i. f i)" by (rule le_SUPI) simp
       
  2250     thus "Real x - (SUP i. f i) \<le> Real x - f i" by (rule pinfreal_minus_le_cancel)
       
  2251   next
       
  2252     fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> y \<le> Real x - f i"
       
  2253     from this[of 0] obtain p where p: "y = Real p" "0 \<le> p"
       
  2254       by (cases "f 0", cases y, auto split: split_if_asm)
       
  2255     hence "\<And>i. Real p \<le> Real x - f i" using * by auto
       
  2256     hence *: "\<And>i. Real x \<le> f i \<Longrightarrow> Real p = 0"
       
  2257       "\<And>i. f i < Real x \<Longrightarrow> Real p + f i \<le> Real x"
       
  2258       unfolding pinfreal_le_minus_iff by auto
       
  2259     show "y \<le> Real x - (SUP i. f i)" unfolding p pinfreal_le_minus_iff
       
  2260     proof safe
       
  2261       assume x_less: "Real x \<le> (SUP i. f i)"
       
  2262       show "Real p = 0"
       
  2263       proof (rule ccontr)
       
  2264         assume "Real p \<noteq> 0"
       
  2265         hence "0 < Real p" by auto
       
  2266         from Sup_close[OF this, of "range f"]
       
  2267         obtain i where e: "(SUP i. f i) < f i + Real p"
       
  2268           using preal unfolding SUPR_def by auto
       
  2269         hence "Real x \<le> f i + Real p" using x_less by auto
       
  2270         show False
       
  2271         proof cases
       
  2272           assume "\<forall>i. f i < Real x"
       
  2273           hence "Real p + f i \<le> Real x" using * by auto
       
  2274           hence "f i + Real p \<le> (SUP i. f i)" using x_less by (auto simp: field_simps)
       
  2275           thus False using e by auto
       
  2276         next
       
  2277           assume "\<not> (\<forall>i. f i < Real x)"
       
  2278           then obtain i where "Real x \<le> f i" by (auto simp: not_less)
       
  2279           from *(1)[OF this] show False using `Real p \<noteq> 0` by auto
       
  2280         qed
       
  2281       qed
       
  2282     next
       
  2283       have "\<And>i. f i \<le> (SUP i. f i)" by (rule complete_lattice_class.le_SUPI) auto
       
  2284       also assume "(SUP i. f i) < Real x"
       
  2285       finally have "\<And>i. f i < Real x" by auto
       
  2286       hence *: "\<And>i. Real p + f i \<le> Real x" using * by auto
       
  2287       have "Real p \<le> Real x" using *[of 0] by (cases "f 0") (auto split: split_if_asm)
       
  2288 
       
  2289       have SUP_eq: "(SUP i. f i) \<le> Real x - Real p"
       
  2290       proof (rule SUP_leI)
       
  2291         fix i show "f i \<le> Real x - Real p" unfolding pinfreal_le_minus_iff
       
  2292         proof safe
       
  2293           assume "Real x \<le> Real p"
       
  2294           with *[of i] show "f i = 0"
       
  2295             by (cases "f i") (auto split: split_if_asm)
       
  2296         next
       
  2297           assume "Real p < Real x"
       
  2298           show "f i + Real p \<le> Real x" using * by (auto simp: field_simps)
       
  2299         qed
       
  2300       qed
       
  2301 
       
  2302       show "Real p + (SUP i. f i) \<le> Real x"
       
  2303       proof cases
       
  2304         assume "Real x \<le> Real p"
       
  2305         with `Real p \<le> Real x` have [simp]: "Real p = Real x" by (rule antisym)
       
  2306         { fix i have "f i = 0" using *[of i] by (cases "f i") (auto split: split_if_asm) }
       
  2307         hence "(SUP i. f i) \<le> 0" by (auto intro!: SUP_leI)
       
  2308         thus ?thesis by simp
       
  2309       next
       
  2310         assume "\<not> Real x \<le> Real p" hence "Real p < Real x" unfolding not_le .
       
  2311         with SUP_eq show ?thesis unfolding pinfreal_le_minus_iff by (auto simp: field_simps)
       
  2312       qed
       
  2313     qed
       
  2314   qed
       
  2315 qed
       
  2316 
       
  2317 lemma pinfreal_SUP_minus:
       
  2318   fixes f :: "nat \<Rightarrow> pinfreal"
       
  2319   shows "(SUP i. c - f i) = c - (INF i. f i)"
       
  2320 proof (rule pinfreal_SUPI)
       
  2321   fix i have "(INF i. f i) \<le> f i" by (rule INF_leI) simp
       
  2322   thus "c - f i \<le> c - (INF i. f i)" by (rule pinfreal_minus_le_cancel)
       
  2323 next
       
  2324   fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c - f i \<le> y"
       
  2325   show "c - (INF i. f i) \<le> y"
       
  2326   proof (cases y)
       
  2327     case (preal p)
       
  2328 
       
  2329     show ?thesis unfolding pinfreal_minus_le_iff preal
       
  2330     proof safe
       
  2331       assume INF_le_x: "(INF i. f i) \<le> c"
       
  2332       from * have *: "\<And>i. f i \<le> c \<Longrightarrow> c \<le> Real p + f i"
       
  2333         unfolding pinfreal_minus_le_iff preal by auto
       
  2334 
       
  2335       have INF_eq: "c - Real p \<le> (INF i. f i)"
       
  2336       proof (rule le_INFI)
       
  2337         fix i show "c - Real p \<le> f i" unfolding pinfreal_minus_le_iff
       
  2338         proof safe
       
  2339           assume "Real p \<le> c"
       
  2340           show "c \<le> f i + Real p"
       
  2341           proof cases
       
  2342             assume "f i \<le> c" from *[OF this]
       
  2343             show ?thesis by (simp add: field_simps)
       
  2344           next
       
  2345             assume "\<not> f i \<le> c"
       
  2346             hence "c \<le> f i" by auto
       
  2347             also have "\<dots> \<le> f i + Real p" by auto
       
  2348             finally show ?thesis .
       
  2349           qed
       
  2350         qed
       
  2351       qed
       
  2352 
       
  2353       show "c \<le> Real p + (INF i. f i)"
       
  2354       proof cases
       
  2355         assume "Real p \<le> c"
       
  2356         with INF_eq show ?thesis unfolding pinfreal_minus_le_iff by (auto simp: field_simps)
       
  2357       next
       
  2358         assume "\<not> Real p \<le> c"
       
  2359         hence "c \<le> Real p" by auto
       
  2360         also have "Real p \<le> Real p + (INF i. f i)" by auto
       
  2361         finally show ?thesis .
       
  2362       qed
       
  2363     qed
       
  2364   qed simp
       
  2365 qed
       
  2366 
       
  2367 lemma pinfreal_le_minus_imp_0:
       
  2368   fixes a b :: pinfreal
       
  2369   shows "a \<le> a - b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a \<noteq> \<omega> \<Longrightarrow> b = 0"
       
  2370   by (cases a, cases b, auto split: split_if_asm)
       
  2371 
       
  2372 lemma lim_INF_le_lim_SUP:
       
  2373   fixes f :: "nat \<Rightarrow> pinfreal"
       
  2374   shows "(SUP n. INF m. f (n + m)) \<le> (INF n. SUP m. f (n + m))"
       
  2375 proof (rule complete_lattice_class.SUP_leI, rule complete_lattice_class.le_INFI)
       
  2376   fix i j show "(INF m. f (i + m)) \<le> (SUP m. f (j + m))"
       
  2377   proof (cases rule: le_cases)
       
  2378     assume "i \<le> j"
       
  2379     have "(INF m. f (i + m)) \<le> f (i + (j - i))" by (rule INF_leI) simp
       
  2380     also have "\<dots> = f (j + 0)" using `i \<le> j` by auto
       
  2381     also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
       
  2382     finally show ?thesis .
       
  2383   next
       
  2384     assume "j \<le> i"
       
  2385     have "(INF m. f (i + m)) \<le> f (i + 0)" by (rule INF_leI) simp
       
  2386     also have "\<dots> = f (j + (i - j))" using `j \<le> i` by auto
       
  2387     also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
       
  2388     finally show ?thesis .
       
  2389   qed
       
  2390 qed
       
  2391 
       
  2392 lemma lim_INF_eq_lim_SUP:
       
  2393   fixes X :: "nat \<Rightarrow> real"
       
  2394   assumes "\<And>i. 0 \<le> X i" and "0 \<le> x"
       
  2395   and lim_INF: "(SUP n. INF m. Real (X (n + m))) = Real x" (is "(SUP n. ?INF n) = _")
       
  2396   and lim_SUP: "(INF n. SUP m. Real (X (n + m))) = Real x" (is "(INF n. ?SUP n) = _")
       
  2397   shows "X ----> x"
       
  2398 proof (rule LIMSEQ_I)
       
  2399   fix r :: real assume "0 < r"
       
  2400   hence "0 \<le> r" by auto
       
  2401   from Sup_close[of "Real r" "range ?INF"]
       
  2402   obtain n where inf: "Real x < ?INF n + Real r"
       
  2403     unfolding SUPR_def lim_INF[unfolded SUPR_def] using `0 < r` by auto
       
  2404 
       
  2405   from Inf_close[of "range ?SUP" "Real r"]
       
  2406   obtain n' where sup: "?SUP n' < Real x + Real r"
       
  2407     unfolding INFI_def lim_SUP[unfolded INFI_def] using `0 < r` by auto
       
  2408 
       
  2409   show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
       
  2410   proof (safe intro!: exI[of _ "max n n'"])
       
  2411     fix m assume "max n n' \<le> m" hence "n \<le> m" "n' \<le> m" by auto
       
  2412 
       
  2413     note inf
       
  2414     also have "?INF n + Real r \<le> Real (X (n + (m - n))) + Real r"
       
  2415       by (rule le_add_Real, auto simp: `0 \<le> r` intro: INF_leI)
       
  2416     finally have up: "x < X m + r"
       
  2417       using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n \<le> m` by auto
       
  2418 
       
  2419     have "Real (X (n' + (m - n'))) \<le> ?SUP n'"
       
  2420       by (auto simp: `0 \<le> r` intro: le_SUPI)
       
  2421     also note sup
       
  2422     finally have down: "X m < x + r"
       
  2423       using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n' \<le> m` by auto
       
  2424 
       
  2425     show "norm (X m - x) < r" using up down by auto
       
  2426   qed
       
  2427 qed
       
  2428 
       
  2429 lemma Sup_countable_SUPR:
       
  2430   assumes "Sup A \<noteq> \<omega>" "A \<noteq> {}"
       
  2431   shows "\<exists> f::nat \<Rightarrow> pinfreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
       
  2432 proof -
       
  2433   have "\<And>n. 0 < 1 / (of_nat n :: pinfreal)" by auto
       
  2434   from Sup_close[OF this assms]
       
  2435   have "\<forall>n. \<exists>x. x \<in> A \<and> Sup A < x + 1 / of_nat n" by blast
       
  2436   from choice[OF this] obtain f where "range f \<subseteq> A" and
       
  2437     epsilon: "\<And>n. Sup A < f n + 1 / of_nat n" by blast
       
  2438   have "SUPR UNIV f = Sup A"
       
  2439   proof (rule pinfreal_SUPI)
       
  2440     fix i show "f i \<le> Sup A" using `range f \<subseteq> A`
       
  2441       by (auto intro!: complete_lattice_class.Sup_upper)
       
  2442   next
       
  2443     fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
       
  2444     show "Sup A \<le> y"
       
  2445     proof (rule pinfreal_le_epsilon)
       
  2446       fix e :: pinfreal assume "0 < e"
       
  2447       show "Sup A \<le> y + e"
       
  2448       proof (cases e)
       
  2449         case (preal r)
       
  2450         hence "0 < r" using `0 < e` by auto
       
  2451         then obtain n where *: "inverse (of_nat n) < r" "0 < n"
       
  2452           using ex_inverse_of_nat_less by auto
       
  2453         have "Sup A \<le> f n + 1 / of_nat n" using epsilon[of n] by auto
       
  2454         also have "1 / of_nat n \<le> e" using preal * by (auto simp: real_eq_of_nat)
       
  2455         with bound have "f n + 1 / of_nat n \<le> y + e" by (rule add_mono) simp
       
  2456         finally show "Sup A \<le> y + e" .
       
  2457       qed simp
       
  2458     qed
       
  2459   qed
       
  2460   with `range f \<subseteq> A` show ?thesis by (auto intro!: exI[of _ f])
       
  2461 qed
       
  2462 
       
  2463 lemma SUPR_countable_SUPR:
       
  2464   assumes "SUPR A g \<noteq> \<omega>" "A \<noteq> {}"
       
  2465   shows "\<exists> f::nat \<Rightarrow> pinfreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
       
  2466 proof -
       
  2467   have "Sup (g`A) \<noteq> \<omega>" "g`A \<noteq> {}" using assms unfolding SUPR_def by auto
       
  2468   from Sup_countable_SUPR[OF this]
       
  2469   show ?thesis unfolding SUPR_def .
       
  2470 qed
       
  2471 
       
  2472 lemma pinfreal_setsum_subtractf:
       
  2473   assumes "\<And>i. i\<in>A \<Longrightarrow> g i \<le> f i" and "\<And>i. i\<in>A \<Longrightarrow> f i \<noteq> \<omega>"
       
  2474   shows "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
       
  2475 proof cases
       
  2476   assume "finite A" from this assms show ?thesis
       
  2477   proof induct
       
  2478     case (insert x A)
       
  2479     hence hyp: "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
       
  2480       by auto
       
  2481     { fix i assume *: "i \<in> insert x A"
       
  2482       hence "g i \<le> f i" using insert by simp
       
  2483       also have "f i < \<omega>" using * insert by (simp add: pinfreal_less_\<omega>)
       
  2484       finally have "g i \<noteq> \<omega>" by (simp add: pinfreal_less_\<omega>) }
       
  2485     hence "setsum g A \<noteq> \<omega>" "g x \<noteq> \<omega>" by (auto simp: setsum_\<omega>)
       
  2486     moreover have "setsum f A \<noteq> \<omega>" "f x \<noteq> \<omega>" using insert by (auto simp: setsum_\<omega>)
       
  2487     moreover have "g x \<le> f x" using insert by auto
       
  2488     moreover have "(\<Sum>i\<in>A. g i) \<le> (\<Sum>i\<in>A. f i)" using insert by (auto intro!: setsum_mono)
       
  2489     ultimately show ?case using `finite A` `x \<notin> A` hyp
       
  2490       by (auto simp: pinfreal_noteq_omega_Ex)
       
  2491   qed simp
       
  2492 qed simp
       
  2493 
       
  2494 lemma real_of_pinfreal_diff:
       
  2495   "y \<le> x \<Longrightarrow> x \<noteq> \<omega> \<Longrightarrow> real x - real y = real (x - y)"
       
  2496   by (cases x, cases y) auto
       
  2497 
       
  2498 lemma psuminf_minus:
       
  2499   assumes ord: "\<And>i. g i \<le> f i" and fin: "psuminf g \<noteq> \<omega>" "psuminf f \<noteq> \<omega>"
       
  2500   shows "(\<Sum>\<^isub>\<infinity> i. f i - g i) = psuminf f - psuminf g"
       
  2501 proof -
       
  2502   have [simp]: "\<And>i. f i \<noteq> \<omega>" using fin by (auto intro: psuminf_\<omega>)
       
  2503   from fin have "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity>x. f x)"
       
  2504     and "(\<lambda>x. real (g x)) sums real (\<Sum>\<^isub>\<infinity>x. g x)"
       
  2505     by (auto intro: psuminf_imp_suminf)
       
  2506   from sums_diff[OF this]
       
  2507   have "(\<lambda>n. real (f n - g n)) sums (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))" using fin ord
       
  2508     by (subst (asm) (1 2) real_of_pinfreal_diff) (auto simp: psuminf_\<omega> psuminf_le)
       
  2509   hence "(\<Sum>\<^isub>\<infinity> i. Real (real (f i - g i))) = Real (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))"
       
  2510     by (rule suminf_imp_psuminf) simp
       
  2511   thus ?thesis using fin by (simp add: Real_real psuminf_\<omega>)
       
  2512 qed
       
  2513 
       
  2514 lemma INF_eq_LIMSEQ:
       
  2515   assumes "mono (\<lambda>i. - f i)" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
       
  2516   shows "(INF n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
       
  2517 proof
       
  2518   assume x: "(INF n. Real (f n)) = Real x"
       
  2519   { fix n
       
  2520     have "Real x \<le> Real (f n)" using x[symmetric] by (auto intro: INF_leI)
       
  2521     hence "x \<le> f n" using assms by simp
       
  2522     hence "\<bar>f n - x\<bar> = f n - x" by auto }
       
  2523   note abs_eq = this
       
  2524   show "f ----> x"
       
  2525   proof (rule LIMSEQ_I)
       
  2526     fix r :: real assume "0 < r"
       
  2527     show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
       
  2528     proof (rule ccontr)
       
  2529       assume *: "\<not> ?thesis"
       
  2530       { fix N
       
  2531         from * obtain n where *: "N \<le> n" "r \<le> f n - x"
       
  2532           using abs_eq by (auto simp: not_less)
       
  2533         hence "x + r \<le> f n" by auto
       
  2534         also have "f n \<le> f N" using `mono (\<lambda>i. - f i)` * by (auto dest: monoD)
       
  2535         finally have "Real (x + r) \<le> Real (f N)" using `0 \<le> f N` by auto }
       
  2536       hence "Real x < Real (x + r)"
       
  2537         and "Real (x + r) \<le> (INF n. Real (f n))" using `0 < r` `0 \<le> x` by (auto intro: le_INFI)
       
  2538       hence "Real x < (INF n. Real (f n))" by (rule less_le_trans)
       
  2539       thus False using x by auto
       
  2540     qed
       
  2541   qed
       
  2542 next
       
  2543   assume "f ----> x"
       
  2544   show "(INF n. Real (f n)) = Real x"
       
  2545   proof (rule pinfreal_INFI)
       
  2546     fix n
       
  2547     from decseq_le[OF _ `f ----> x`] assms
       
  2548     show "Real x \<le> Real (f n)" unfolding decseq_eq_incseq incseq_mono by auto
       
  2549   next
       
  2550     fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> y \<le> Real (f n)"
       
  2551     thus "y \<le> Real x"
       
  2552     proof (cases y)
       
  2553       case (preal r)
       
  2554       with * have "\<exists>N. \<forall>n\<ge>N. r \<le> f n" using assms by fastsimp
       
  2555       from LIMSEQ_le_const[OF `f ----> x` this]
       
  2556       show "y \<le> Real x" using `0 \<le> x` preal by auto
       
  2557     qed simp
       
  2558   qed
       
  2559 qed
       
  2560 
       
  2561 lemma INFI_bound:
       
  2562   assumes "\<forall>N. x \<le> f N"
       
  2563   shows "x \<le> (INF n. f n)"
       
  2564   using assms by (simp add: INFI_def le_Inf_iff)
       
  2565 
       
  2566 lemma INF_mono:
       
  2567   assumes "\<And>n. f (N n) \<le> g n"
       
  2568   shows "(INF n. f n) \<le> (INF n. g n)"
       
  2569 proof (safe intro!: INFI_bound)
       
  2570   fix n
       
  2571   have "(INF n. f n) \<le> f (N n)" by (auto intro!: INF_leI)
       
  2572   also note assms[of n]
       
  2573   finally show "(INF n. f n) \<le> g n" .
       
  2574 qed
       
  2575 
       
  2576 lemma INFI_fun_expand: "(INF y:A. f y) = (\<lambda>x. INF y:A. f y x)"
       
  2577   unfolding INFI_def expand_fun_eq Inf_fun_def
       
  2578   by (auto intro!: arg_cong[where f=Inf])
       
  2579 
       
  2580 lemma LIMSEQ_imp_lim_INF:
       
  2581   assumes pos: "\<And>i. 0 \<le> X i" and lim: "X ----> x"
       
  2582   shows "(SUP n. INF m. Real (X (n + m))) = Real x"
       
  2583 proof -
       
  2584   have "0 \<le> x" using assms by (auto intro!: LIMSEQ_le_const)
       
  2585 
       
  2586   have "\<And>n. (INF m. Real (X (n + m))) \<le> Real (X (n + 0))" by (rule INF_leI) simp
       
  2587   also have "\<And>n. Real (X (n + 0)) < \<omega>" by simp
       
  2588   finally have "\<forall>n. \<exists>r\<ge>0. (INF m. Real (X (n + m))) = Real r"
       
  2589     by (auto simp: pinfreal_less_\<omega> pinfreal_noteq_omega_Ex)
       
  2590   from choice[OF this] obtain r where r: "\<And>n. (INF m. Real (X (n + m))) = Real (r n)" "\<And>n. 0 \<le> r n"
       
  2591     by auto
       
  2592 
       
  2593   show ?thesis unfolding r
       
  2594   proof (subst SUP_eq_LIMSEQ)
       
  2595     show "mono r" unfolding mono_def
       
  2596     proof safe
       
  2597       fix x y :: nat assume "x \<le> y"
       
  2598       have "Real (r x) \<le> Real (r y)" unfolding r(1)[symmetric] using pos
       
  2599       proof (safe intro!: INF_mono)
       
  2600         fix m have "x + (m + y - x) = y + m"
       
  2601           using `x \<le> y` by auto
       
  2602         thus "Real (X (x + (m + y - x))) \<le> Real (X (y + m))" by simp
       
  2603       qed
       
  2604       thus "r x \<le> r y" using r by auto
       
  2605     qed
       
  2606     show "\<And>n. 0 \<le> r n" by fact
       
  2607     show "0 \<le> x" by fact
       
  2608     show "r ----> x"
       
  2609     proof (rule LIMSEQ_I)
       
  2610       fix e :: real assume "0 < e"
       
  2611       hence "0 < e/2" by auto
       
  2612       from LIMSEQ_D[OF lim this] obtain N where *: "\<And>n. N \<le> n \<Longrightarrow> \<bar>X n - x\<bar> < e/2"
       
  2613         by auto
       
  2614       show "\<exists>N. \<forall>n\<ge>N. norm (r n - x) < e"
       
  2615       proof (safe intro!: exI[of _ N])
       
  2616         fix n assume "N \<le> n"
       
  2617         show "norm (r n - x) < e"
       
  2618         proof cases
       
  2619           assume "r n < x"
       
  2620           have "x - r n \<le> e/2"
       
  2621           proof cases
       
  2622             assume e: "e/2 \<le> x"
       
  2623             have "Real (x - e/2) \<le> Real (r n)" unfolding r(1)[symmetric]
       
  2624             proof (rule le_INFI)
       
  2625               fix m show "Real (x - e / 2) \<le> Real (X (n + m))"
       
  2626                 using *[of "n + m"] `N \<le> n`
       
  2627                 using pos by (auto simp: field_simps abs_real_def split: split_if_asm)
       
  2628             qed
       
  2629             with e show ?thesis using pos `0 \<le> x` r(2) by auto
       
  2630           next
       
  2631             assume "\<not> e/2 \<le> x" hence "x - e/2 < 0" by auto
       
  2632             with `0 \<le> r n` show ?thesis by auto
       
  2633           qed
       
  2634           with `r n < x` show ?thesis by simp
       
  2635         next
       
  2636           assume e: "\<not> r n < x"
       
  2637           have "Real (r n) \<le> Real (X (n + 0))" unfolding r(1)[symmetric]
       
  2638             by (rule INF_leI) simp
       
  2639           hence "r n - x \<le> X n - x" using r pos by auto
       
  2640           also have "\<dots> < e/2" using *[OF `N \<le> n`] by (auto simp: field_simps abs_real_def split: split_if_asm)
       
  2641           finally have "r n - x < e" using `0 < e` by auto
       
  2642           with e show ?thesis by auto
       
  2643         qed
       
  2644       qed
       
  2645     qed
       
  2646   qed
       
  2647 qed
       
  2648 
       
  2649 
       
  2650 lemma real_of_pinfreal_strict_mono_iff:
       
  2651   "real a < real b \<longleftrightarrow> (b \<noteq> \<omega> \<and> ((a = \<omega> \<and> 0 < b) \<or> (a < b)))"
       
  2652 proof (cases a)
       
  2653   case infinite thus ?thesis by (cases b) auto
       
  2654 next
       
  2655   case preal thus ?thesis by (cases b) auto
       
  2656 qed
       
  2657 
       
  2658 lemma real_of_pinfreal_mono_iff:
       
  2659   "real a \<le> real b \<longleftrightarrow> (a = \<omega> \<or> (b \<noteq> \<omega> \<and> a \<le> b) \<or> (b = \<omega> \<and> a = 0))"
       
  2660 proof (cases a)
       
  2661   case infinite thus ?thesis by (cases b) auto
       
  2662 next
       
  2663   case preal thus ?thesis by (cases b)  auto
       
  2664 qed
       
  2665 
       
  2666 lemma ex_pinfreal_inverse_of_nat_Suc_less:
       
  2667   fixes e :: pinfreal assumes "0 < e" shows "\<exists>n. inverse (of_nat (Suc n)) < e"
       
  2668 proof (cases e)
       
  2669   case (preal r)
       
  2670   with `0 < e` ex_inverse_of_nat_Suc_less[of r]
       
  2671   obtain n where "inverse (of_nat (Suc n)) < r" by auto
       
  2672   with preal show ?thesis
       
  2673     by (auto simp: real_eq_of_nat[symmetric])
       
  2674 qed auto
       
  2675 
       
  2676 lemma Lim_eq_Sup_mono:
       
  2677   fixes u :: "nat \<Rightarrow> pinfreal" assumes "mono u"
       
  2678   shows "u ----> (SUP i. u i)"
       
  2679 proof -
       
  2680   from lim_pinfreal_increasing[of u] `mono u`
       
  2681   obtain l where l: "u ----> l" unfolding mono_def by auto
       
  2682   from SUP_Lim_pinfreal[OF _ this] `mono u`
       
  2683   have "(SUP i. u i) = l" unfolding mono_def by auto
       
  2684   with l show ?thesis by simp
       
  2685 qed
       
  2686 
       
  2687 lemma isotone_Lim:
       
  2688   fixes x :: pinfreal assumes "u \<up> x"
       
  2689   shows "u ----> x" (is ?lim) and "mono u" (is ?mono)
       
  2690 proof -
       
  2691   show ?mono using assms unfolding mono_iff_le_Suc isoton_def by auto
       
  2692   from Lim_eq_Sup_mono[OF this] `u \<up> x`
       
  2693   show ?lim unfolding isoton_def by simp
       
  2694 qed
       
  2695 
       
  2696 lemma isoton_iff_Lim_mono:
       
  2697   fixes u :: "nat \<Rightarrow> pinfreal"
       
  2698   shows "u \<up> x \<longleftrightarrow> (mono u \<and> u ----> x)"
       
  2699 proof safe
       
  2700   assume "mono u" and x: "u ----> x"
       
  2701   with SUP_Lim_pinfreal[OF _ x]
       
  2702   show "u \<up> x" unfolding isoton_def
       
  2703     using `mono u`[unfolded mono_def]
       
  2704     using `mono u`[unfolded mono_iff_le_Suc]
       
  2705     by auto
       
  2706 qed (auto dest: isotone_Lim)
       
  2707 
       
  2708 lemma pinfreal_inverse_inverse[simp]:
       
  2709   fixes x :: pinfreal
       
  2710   shows "inverse (inverse x) = x"
       
  2711   by (cases x) auto
       
  2712 
       
  2713 lemma atLeastAtMost_omega_eq_atLeast:
       
  2714   shows "{a .. \<omega>} = {a ..}"
       
  2715 by auto
       
  2716 
       
  2717 lemma atLeast0AtMost_eq_atMost: "{0 :: pinfreal .. a} = {.. a}" by auto
       
  2718 
       
  2719 lemma greaterThan_omega_Empty: "{\<omega> <..} = {}" by auto
       
  2720 
       
  2721 lemma lessThan_0_Empty: "{..< 0 :: pinfreal} = {}" by auto
       
  2722 
       
  2723 end