| author | hoelzl | 
| Fri, 16 Nov 2012 18:45:57 +0100 | |
| changeset 50104 | de19856feb54 | 
| parent 49664 | f099b8006a3c | 
| child 51000 | c9adb50f74ad | 
| permissions | -rw-r--r-- | 
| 41983 | 1 | (* Title: HOL/Multivariate_Analysis/Extended_Real_Limits.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | Author: Robert Himmelmann, TU München | |
| 4 | Author: Armin Heller, TU München | |
| 5 | Author: Bogdan Grechuk, University of Edinburgh | |
| 6 | *) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 7 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 8 | header {* Limits on the Extended real number line *}
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 9 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 10 | theory Extended_Real_Limits | 
| 43920 | 11 | imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Real" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 12 | begin | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 13 | |
| 43920 | 14 | lemma continuous_on_ereal[intro, simp]: "continuous_on A ereal" | 
| 15 | unfolding continuous_on_topological open_ereal_def by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 16 | |
| 43920 | 17 | lemma continuous_at_ereal[intro, simp]: "continuous (at x) ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 18 | using continuous_on_eq_continuous_at[of UNIV] by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 19 | |
| 43920 | 20 | lemma continuous_within_ereal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 21 | using continuous_on_eq_continuous_within[of A] by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 22 | |
| 43920 | 23 | lemma ereal_open_uminus: | 
| 24 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 25 | assumes "open S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 26 | shows "open (uminus ` S)" | 
| 43920 | 27 | unfolding open_ereal_def | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 28 | proof (intro conjI impI) | 
| 49664 | 29 | obtain x y where | 
| 30 |     S: "open (ereal -` S)" "\<infinity> \<in> S \<Longrightarrow> {ereal x<..} \<subseteq> S" "-\<infinity> \<in> S \<Longrightarrow> {..< ereal y} \<subseteq> S"
 | |
| 43920 | 31 | using `open S` unfolding open_ereal_def by auto | 
| 32 | have "ereal -` uminus ` S = uminus ` (ereal -` S)" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 33 | proof safe | 
| 49664 | 34 | fix x y | 
| 35 | assume "ereal x = - y" "y \<in> S" | |
| 43920 | 36 | then show "x \<in> uminus ` ereal -` S" by (cases y) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 37 | next | 
| 49664 | 38 | fix x | 
| 39 | assume "ereal x \<in> S" | |
| 43920 | 40 | then show "- x \<in> ereal -` uminus ` S" | 
| 41 | by (auto intro: image_eqI[of _ _ "ereal x"]) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 42 | qed | 
| 43920 | 43 | then show "open (ereal -` uminus ` S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 44 | using S by (auto intro: open_negations) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 45 |   { assume "\<infinity> \<in> uminus ` S"
 | 
| 43920 | 46 | then have "-\<infinity> \<in> S" by (metis image_iff ereal_uminus_uminus) | 
| 47 |     then have "uminus ` {..<ereal y} \<subseteq> uminus ` S" using S by (intro image_mono) auto
 | |
| 48 |     then show "\<exists>x. {ereal x<..} \<subseteq> uminus ` S" using ereal_uminus_lessThan by auto }
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 49 |   { assume "-\<infinity> \<in> uminus ` S"
 | 
| 43920 | 50 | then have "\<infinity> : S" by (metis image_iff ereal_uminus_uminus) | 
| 51 |     then have "uminus ` {ereal x<..} <= uminus ` S" using S by (intro image_mono) auto
 | |
| 52 |     then show "\<exists>y. {..<ereal y} <= uminus ` S" using ereal_uminus_greaterThan by auto }
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 53 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 54 | |
| 43920 | 55 | lemma ereal_uminus_complement: | 
| 56 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 57 | shows "uminus ` (- S) = - uminus ` S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 58 | by (auto intro!: bij_image_Compl_eq surjI[of _ uminus] simp: bij_betw_def) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 59 | |
| 43920 | 60 | lemma ereal_closed_uminus: | 
| 61 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 62 | assumes "closed S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 63 | shows "closed (uminus ` S)" | 
| 49664 | 64 | using assms unfolding closed_def | 
| 65 | using ereal_open_uminus[of "- S"] ereal_uminus_complement by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 66 | |
| 44571 | 67 | instance ereal :: perfect_space | 
| 68 | proof (default, rule) | |
| 69 |   fix a :: ereal assume a: "open {a}"
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 70 | show False | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 71 | proof (cases a) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 72 | case MInf | 
| 43920 | 73 |     then obtain y where "{..<ereal y} <= {a}" using a open_MInfty2[of "{a}"] by auto
 | 
| 49664 | 74 |     then have "ereal(y - 1):{a}" apply (subst subsetD[of "{..<ereal y}"]) by auto
 | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 75 | then show False using `a=(-\<infinity>)` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 76 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 77 | case PInf | 
| 43920 | 78 |     then obtain y where "{ereal y<..} <= {a}" using a open_PInfty2[of "{a}"] by auto
 | 
| 49664 | 79 |     then have "ereal(y+1):{a}" apply (subst subsetD[of "{ereal y<..}"]) by auto
 | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 80 | then show False using `a=\<infinity>` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 81 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 82 | case (real r) then have fin: "\<bar>a\<bar> \<noteq> \<infinity>" by simp | 
| 43920 | 83 | from ereal_open_cont_interval[OF a singletonI this] guess e . note e = this | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 84 | then obtain b where b_def: "a<b & b<a+e" | 
| 43920 | 85 | using fin ereal_between ereal_dense[of a "a+e"] by auto | 
| 86 |     then have "b: {a-e <..< a+e}" using fin ereal_between[of a e] e by auto
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 87 | then show False using b_def e by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 88 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 89 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 90 | |
| 43920 | 91 | lemma ereal_closed_contains_Inf: | 
| 92 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 93 |   assumes "closed S" "S ~= {}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 94 | shows "Inf S : S" | 
| 49664 | 95 | proof (rule ccontr) | 
| 96 | assume "Inf S \<notin> S" | |
| 97 | then have a: "open (-S)" "Inf S:(- S)" using assms by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 98 | show False | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 99 | proof (cases "Inf S") | 
| 49664 | 100 | case MInf | 
| 101 | then have "(-\<infinity>) : - S" using a by auto | |
| 43920 | 102 |     then obtain y where "{..<ereal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
 | 
| 49664 | 103 | then have "ereal y <= Inf S" by (metis Compl_anti_mono Compl_lessThan atLeast_iff | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 104 | complete_lattice_class.Inf_greatest double_complement set_rev_mp) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 105 | then show False using MInf by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 106 | next | 
| 49664 | 107 | case PInf | 
| 108 |     then have "S={\<infinity>}" by (metis Inf_eq_PInfty assms(2))
 | |
| 44918 | 109 | then show False using `Inf S ~: S` by (simp add: top_ereal_def) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 110 | next | 
| 49664 | 111 | case (real r) | 
| 112 | then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp | |
| 43920 | 113 | from ereal_open_cont_interval[OF a this] guess e . note e = this | 
| 49664 | 114 |     { fix x
 | 
| 115 | assume "x:S" then have "x>=Inf S" by (rule complete_lattice_class.Inf_lower) | |
| 116 | then have *: "x>Inf S-e" using e by (metis fin ereal_between(1) order_less_le_trans) | |
| 117 |       { assume "x<Inf S+e"
 | |
| 118 |         then have "x:{Inf S-e <..< Inf S+e}" using * by auto
 | |
| 119 | then have False using e `x:S` by auto | |
| 120 | } then have "x>=Inf S+e" by (metis linorder_le_less_linear) | |
| 121 | } | |
| 122 | then have "Inf S + e <= Inf S" by (metis le_Inf_iff) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 123 | then show False using real e by (cases e) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 124 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 125 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 126 | |
| 43920 | 127 | lemma ereal_closed_contains_Sup: | 
| 128 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 129 |   assumes "closed S" "S ~= {}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 130 | shows "Sup S : S" | 
| 49664 | 131 | proof - | 
| 132 | have "closed (uminus ` S)" | |
| 133 | by (metis assms(1) ereal_closed_uminus) | |
| 134 | then have "Inf (uminus ` S) : uminus ` S" | |
| 135 | using assms ereal_closed_contains_Inf[of "uminus ` S"] by auto | |
| 136 | then have "- Sup S : uminus ` S" | |
| 137 | using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image) | |
| 138 | then show ?thesis | |
| 139 | by (metis imageI ereal_uminus_uminus ereal_minus_minus_image) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 140 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 141 | |
| 43920 | 142 | lemma ereal_open_closed_aux: | 
| 143 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 144 | assumes "open S" "closed S" | 
| 49664 | 145 | and S: "(-\<infinity>) ~: S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 146 |   shows "S = {}"
 | 
| 49664 | 147 | proof (rule ccontr) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 148 |   assume "S ~= {}"
 | 
| 49664 | 149 | then have *: "(Inf S):S" by (metis assms(2) ereal_closed_contains_Inf) | 
| 150 |   { assume "Inf S=(-\<infinity>)"
 | |
| 151 | then have False using * assms(3) by auto } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 152 | moreover | 
| 49664 | 153 |   { assume "Inf S=\<infinity>"
 | 
| 154 |     then have "S={\<infinity>}" by (metis Inf_eq_PInfty `S ~= {}`)
 | |
| 155 | then have False by (metis assms(1) not_open_singleton) } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 156 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 157 |   { assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"
 | 
| 43920 | 158 | from ereal_open_cont_interval[OF assms(1) * fin] guess e . note e = this | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 159 | then obtain b where b_def: "Inf S-e<b & b<Inf S" | 
| 43920 | 160 | using fin ereal_between[of "Inf S" e] ereal_dense[of "Inf S-e"] by auto | 
| 49664 | 161 |     then have "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e]
 | 
| 44918 | 162 | by auto | 
| 49664 | 163 | then have "b:S" using e by auto | 
| 164 | then have False using b_def by (metis complete_lattice_class.Inf_lower leD) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 165 | } ultimately show False by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 166 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 167 | |
| 43920 | 168 | lemma ereal_open_closed: | 
| 169 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 170 |   shows "(open S & closed S) <-> (S = {} | S = UNIV)"
 | 
| 49664 | 171 | proof - | 
| 172 |   { assume lhs: "open S & closed S"
 | |
| 173 |     { assume "(-\<infinity>) ~: S"
 | |
| 174 |       then have "S={}" using lhs ereal_open_closed_aux by auto }
 | |
| 175 | moreover | |
| 176 |     { assume "(-\<infinity>) : S"
 | |
| 177 |       then have "(- S)={}" using lhs ereal_open_closed_aux[of "-S"] by auto }
 | |
| 178 |     ultimately have "S = {} | S = UNIV" by auto
 | |
| 179 | } then show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 180 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 181 | |
| 43920 | 182 | lemma ereal_open_affinity_pos: | 
| 43923 | 183 | fixes S :: "ereal set" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 184 | assumes "open S" and m: "m \<noteq> \<infinity>" "0 < m" and t: "\<bar>t\<bar> \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 185 | shows "open ((\<lambda>x. m * x + t) ` S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 186 | proof - | 
| 43920 | 187 | obtain r where r[simp]: "m = ereal r" using m by (cases m) auto | 
| 188 | obtain p where p[simp]: "t = ereal p" using t by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 189 | have "r \<noteq> 0" "0 < r" and m': "m \<noteq> \<infinity>" "m \<noteq> -\<infinity>" "m \<noteq> 0" using m by auto | 
| 43920 | 190 | from `open S`[THEN ereal_openE] guess l u . note T = this | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 191 | let ?f = "(\<lambda>x. m * x + t)" | 
| 49664 | 192 | show ?thesis | 
| 193 | unfolding open_ereal_def | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 194 | proof (intro conjI impI exI subsetI) | 
| 43920 | 195 | have "ereal -` ?f ` S = (\<lambda>x. r * x + p) ` (ereal -` S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 196 | proof safe | 
| 49664 | 197 | fix x y | 
| 198 | assume "ereal y = m * x + t" "x \<in> S" | |
| 43920 | 199 | then show "y \<in> (\<lambda>x. r * x + p) ` ereal -` S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 200 | using `r \<noteq> 0` by (cases x) (auto intro!: image_eqI[of _ _ "real x"] split: split_if_asm) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 201 | qed force | 
| 43920 | 202 | then show "open (ereal -` ?f ` S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 203 | using open_affinity[OF T(1) `r \<noteq> 0`] by (auto simp: ac_simps) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 204 | next | 
| 49664 | 205 | assume "\<infinity> \<in> ?f`S" | 
| 206 | with `0 < r` have "\<infinity> \<in> S" by auto | |
| 207 | fix x | |
| 208 |     assume "x \<in> {ereal (r * l + p)<..}"
 | |
| 43920 | 209 | then have [simp]: "ereal (r * l + p) < x" by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 210 | show "x \<in> ?f`S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 211 | proof (rule image_eqI) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 212 | show "x = m * ((x - t) / m) + t" | 
| 43920 | 213 | using m t by (cases rule: ereal3_cases[of m x t]) auto | 
| 214 | have "ereal l < (x - t)/m" | |
| 215 | using m t by (simp add: ereal_less_divide_pos ereal_less_minus) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 216 | then show "(x - t)/m \<in> S" using T(2)[OF `\<infinity> \<in> S`] by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 217 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 218 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 219 | assume "-\<infinity> \<in> ?f`S" with `0 < r` have "-\<infinity> \<in> S" by auto | 
| 43920 | 220 |     fix x assume "x \<in> {..<ereal (r * u + p)}"
 | 
| 221 | then have [simp]: "x < ereal (r * u + p)" by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 222 | show "x \<in> ?f`S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 223 | proof (rule image_eqI) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 224 | show "x = m * ((x - t) / m) + t" | 
| 43920 | 225 | using m t by (cases rule: ereal3_cases[of m x t]) auto | 
| 226 | have "(x - t)/m < ereal u" | |
| 227 | using m t by (simp add: ereal_divide_less_pos ereal_minus_less) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 228 | then show "(x - t)/m \<in> S" using T(3)[OF `-\<infinity> \<in> S`] by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 229 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 230 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 231 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 232 | |
| 43920 | 233 | lemma ereal_open_affinity: | 
| 43923 | 234 | fixes S :: "ereal set" | 
| 49664 | 235 | assumes "open S" | 
| 236 | and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" | |
| 237 | and t: "\<bar>t\<bar> \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 238 | shows "open ((\<lambda>x. m * x + t) ` S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 239 | proof cases | 
| 49664 | 240 | assume "0 < m" | 
| 241 | then show ?thesis | |
| 43920 | 242 | using ereal_open_affinity_pos[OF `open S` _ _ t, of m] m by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 243 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 244 | assume "\<not> 0 < m" then | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 245 | have "0 < -m" using `m \<noteq> 0` by (cases m) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 246 | then have m: "-m \<noteq> \<infinity>" "0 < -m" using `\<bar>m\<bar> \<noteq> \<infinity>` | 
| 43920 | 247 | by (auto simp: ereal_uminus_eq_reorder) | 
| 248 | from ereal_open_affinity_pos[OF ereal_open_uminus[OF `open S`] m t] | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 249 | show ?thesis unfolding image_image by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 250 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 251 | |
| 43920 | 252 | lemma ereal_lim_mult: | 
| 253 | fixes X :: "'a \<Rightarrow> ereal" | |
| 49664 | 254 | assumes lim: "(X ---> L) net" | 
| 255 | and a: "\<bar>a\<bar> \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 256 | shows "((\<lambda>i. a * X i) ---> a * L) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 257 | proof cases | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 258 | assume "a \<noteq> 0" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 259 | show ?thesis | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 260 | proof (rule topological_tendstoI) | 
| 49664 | 261 | fix S | 
| 262 | assume "open S" "a * L \<in> S" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 263 | have "a * L / a = L" | 
| 43920 | 264 | using `a \<noteq> 0` a by (cases rule: ereal2_cases[of a L]) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 265 | then have L: "L \<in> ((\<lambda>x. x / a) ` S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 266 | using `a * L \<in> S` by (force simp: image_iff) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 267 | moreover have "open ((\<lambda>x. x / a) ` S)" | 
| 43920 | 268 | using ereal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a | 
| 269 | by (auto simp: ereal_divide_eq ereal_inverse_eq_0 divide_ereal_def ac_simps) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 270 | note * = lim[THEN topological_tendstoD, OF this L] | 
| 49664 | 271 |     { fix x
 | 
| 272 | from a `a \<noteq> 0` have "a * (x / a) = x" | |
| 43920 | 273 | by (cases rule: ereal2_cases[of a x]) auto } | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 274 | note this[simp] | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 275 | show "eventually (\<lambda>x. a * X x \<in> S) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 276 | by (rule eventually_mono[OF _ *]) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 277 | qed | 
| 44918 | 278 | qed auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 279 | |
| 43920 | 280 | lemma ereal_lim_uminus: | 
| 49664 | 281 | fixes X :: "'a \<Rightarrow> ereal" | 
| 282 | shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net" | |
| 43920 | 283 | using ereal_lim_mult[of X L net "ereal (-1)"] | 
| 49664 | 284 | ereal_lim_mult[of "(\<lambda>i. - X i)" "-L" net "ereal (-1)"] | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 285 | by (auto simp add: algebra_simps) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 286 | |
| 43920 | 287 | lemma Lim_bounded2_ereal: | 
| 288 | assumes lim:"f ----> (l :: ereal)" | |
| 49664 | 289 | and ge: "ALL n>=N. f n >= C" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 290 | shows "l>=C" | 
| 49664 | 291 | proof - | 
| 292 | def g == "(%i. -(f i))" | |
| 293 |   { fix n
 | |
| 294 | assume "n>=N" | |
| 295 | then have "g n <= -C" using assms ereal_minus_le_minus g_def by auto } | |
| 296 | then have "ALL n>=N. g n <= -C" by auto | |
| 297 | moreover have limg: "g ----> (-l)" using g_def ereal_lim_uminus lim by auto | |
| 298 | ultimately have "-l <= -C" using Lim_bounded_ereal[of g "-l" _ "-C"] by auto | |
| 299 | then show ?thesis using ereal_minus_le_minus by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 300 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 301 | |
| 43923 | 302 | lemma ereal_open_atLeast: fixes x :: ereal shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
 | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 303 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 304 |   assume "x = -\<infinity>" then have "{x..} = UNIV" by auto
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 305 |   then show "open {x..}" by auto
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 306 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 307 |   assume "open {x..}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 308 |   then have "open {x..} \<and> closed {x..}" by auto
 | 
| 43920 | 309 |   then have "{x..} = UNIV" unfolding ereal_open_closed by auto
 | 
| 310 | then show "x = -\<infinity>" by (simp add: bot_ereal_def atLeast_eq_UNIV_iff) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 311 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 312 | |
| 43920 | 313 | lemma ereal_open_mono_set: | 
| 314 | fixes S :: "ereal set" | |
| 49664 | 315 |   shows "(open S \<and> mono_set S) \<longleftrightarrow> (S = UNIV \<or> S = {Inf S <..})"
 | 
| 316 | by (metis Inf_UNIV atLeast_eq_UNIV_iff ereal_open_atLeast | |
| 317 | ereal_open_closed mono_set_iff open_ereal_greaterThan) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 318 | |
| 43920 | 319 | lemma ereal_closed_mono_set: | 
| 320 | fixes S :: "ereal set" | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44167diff
changeset | 321 |   shows "(closed S \<and> mono_set S) \<longleftrightarrow> (S = {} \<or> S = {Inf S ..})"
 | 
| 43920 | 322 | by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast | 
| 49664 | 323 | ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 324 | |
| 43920 | 325 | lemma ereal_Liminf_Sup_monoset: | 
| 326 | fixes f :: "'a => ereal" | |
| 49664 | 327 | shows "Liminf net f = | 
| 328 |     Sup {l. \<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 329 | unfolding Liminf_Sup | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 330 | proof (intro arg_cong[where f="\<lambda>P. Sup (Collect P)"] ext iffI allI impI) | 
| 49664 | 331 | fix l S | 
| 332 | assume ev: "\<forall>y<l. eventually (\<lambda>x. y < f x) net" and "open S" "mono_set S" "l \<in> S" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 333 |   then have "S = UNIV \<or> S = {Inf S <..}"
 | 
| 43920 | 334 | using ereal_open_mono_set[of S] by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 335 | then show "eventually (\<lambda>x. f x \<in> S) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 336 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 337 |     assume S: "S = {Inf S<..}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 338 | then have "Inf S < l" using `l \<in> S` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 339 | then have "eventually (\<lambda>x. Inf S < f x) net" using ev by auto | 
| 44918 | 340 | then show "eventually (\<lambda>x. f x \<in> S) net" by (subst S) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 341 | qed auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 342 | next | 
| 49664 | 343 | fix l y | 
| 344 | assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" "y < l" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 345 |   have "eventually  (\<lambda>x. f x \<in> {y <..}) net"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 346 | using `y < l` by (intro S[rule_format]) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 347 | then show "eventually (\<lambda>x. y < f x) net" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 348 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 349 | |
| 43920 | 350 | lemma ereal_Limsup_Inf_monoset: | 
| 351 | fixes f :: "'a => ereal" | |
| 49664 | 352 | shows "Limsup net f = | 
| 353 |     Inf {l. \<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 354 | unfolding Limsup_Inf | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 355 | proof (intro arg_cong[where f="\<lambda>P. Inf (Collect P)"] ext iffI allI impI) | 
| 49664 | 356 | fix l S | 
| 357 | assume ev: "\<forall>y>l. eventually (\<lambda>x. f x < y) net" and "open S" "mono_set (uminus`S)" "l \<in> S" | |
| 44170 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 huffman parents: 
44167diff
changeset | 358 | then have "open (uminus`S) \<and> mono_set (uminus`S)" by (simp add: ereal_open_uminus) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 359 |   then have "S = UNIV \<or> S = {..< Sup S}"
 | 
| 43920 | 360 | unfolding ereal_open_mono_set ereal_Inf_uminus_image_eq ereal_image_uminus_shift by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 361 | then show "eventually (\<lambda>x. f x \<in> S) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 362 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 363 |     assume S: "S = {..< Sup S}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 364 | then have "l < Sup S" using `l \<in> S` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 365 | then have "eventually (\<lambda>x. f x < Sup S) net" using ev by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 366 | then show "eventually (\<lambda>x. f x \<in> S) net" by (subst S) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 367 | qed auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 368 | next | 
| 49664 | 369 | fix l y | 
| 370 | assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus`S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" "l < y" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 371 |   have "eventually  (\<lambda>x. f x \<in> {..< y}) net"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 372 | using `l < y` by (intro S[rule_format]) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 373 | then show "eventually (\<lambda>x. f x < y) net" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 374 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 375 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 376 | |
| 43920 | 377 | lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::ereal set)" | 
| 378 | using ereal_open_uminus[of S] ereal_open_uminus[of "uminus`S"] by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 379 | |
| 43920 | 380 | lemma ereal_Limsup_uminus: | 
| 381 | fixes f :: "'a => ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 382 | shows "Limsup net (\<lambda>x. - (f x)) = -(Liminf net f)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 383 | proof - | 
| 49664 | 384 |   { fix P l
 | 
| 385 | have "(\<exists>x. (l::ereal) = -x \<and> P x) \<longleftrightarrow> P (-l)" | |
| 386 | by (auto intro!: exI[of _ "-l"]) } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 387 | note Ex_cancel = this | 
| 49664 | 388 |   { fix P :: "ereal set \<Rightarrow> bool"
 | 
| 389 | have "(\<forall>S. P S) \<longleftrightarrow> (\<forall>S. P (uminus`S))" | |
| 390 | apply auto | |
| 391 | apply (erule_tac x="uminus`S" in allE) | |
| 392 | apply (auto simp: image_image) | |
| 393 | done } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 394 | note add_uminus_image = this | 
| 49664 | 395 |   { fix x S
 | 
| 396 | have "(x::ereal) \<in> uminus`S \<longleftrightarrow> -x\<in>S" | |
| 397 | by (auto intro!: image_eqI[of _ _ "-x"]) } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 398 | note remove_uminus_image = this | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 399 | show ?thesis | 
| 43920 | 400 | unfolding ereal_Limsup_Inf_monoset ereal_Liminf_Sup_monoset | 
| 401 | unfolding ereal_Inf_uminus_image_eq[symmetric] image_Collect Ex_cancel | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 402 | by (subst add_uminus_image) (simp add: open_uminus_iff remove_uminus_image) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 403 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 404 | |
| 43920 | 405 | lemma ereal_Liminf_uminus: | 
| 406 | fixes f :: "'a => ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 407 | shows "Liminf net (\<lambda>x. - (f x)) = -(Limsup net f)" | 
| 43920 | 408 | using ereal_Limsup_uminus[of _ "(\<lambda>x. - (f x))"] by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 409 | |
| 43920 | 410 | lemma ereal_Lim_uminus: | 
| 49664 | 411 | fixes f :: "'a \<Rightarrow> ereal" | 
| 412 | shows "(f ---> f0) net \<longleftrightarrow> ((\<lambda>x. - f x) ---> - f0) net" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 413 | using | 
| 43920 | 414 | ereal_lim_mult[of f f0 net "- 1"] | 
| 415 | ereal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"] | |
| 416 | by (auto simp: ereal_uminus_reorder) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 417 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 418 | lemma lim_imp_Limsup: | 
| 43920 | 419 | fixes f :: "'a => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 420 | assumes "\<not> trivial_limit net" | 
| 49664 | 421 | and lim: "(f ---> f0) net" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 422 | shows "Limsup net f = f0" | 
| 43920 | 423 | using ereal_Lim_uminus[of f f0] lim_imp_Liminf[of net "(%x. -(f x))" "-f0"] | 
| 424 | ereal_Liminf_uminus[of net f] assms by simp | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 425 | |
| 50104 | 426 | lemma convergent_ereal_limsup: | 
| 427 | fixes X :: "nat \<Rightarrow> ereal" | |
| 428 | shows "convergent X \<Longrightarrow> limsup X = lim X" | |
| 429 | by (auto simp: convergent_def limI lim_imp_Limsup) | |
| 430 | ||
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 431 | lemma Liminf_PInfty: | 
| 43920 | 432 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 433 | assumes "\<not> trivial_limit net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 434 | shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 435 | proof (intro lim_imp_Liminf iffI assms) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 436 | assume rhs: "Liminf net f = \<infinity>" | 
| 49664 | 437 |   { fix S :: "ereal set"
 | 
| 438 | assume "open S & \<infinity> : S" | |
| 43920 | 439 |     then obtain m where "{ereal m<..} <= S" using open_PInfty2 by auto
 | 
| 49664 | 440 | moreover | 
| 441 |     have "eventually (\<lambda>x. f x \<in> {ereal m<..}) net"
 | |
| 442 | using rhs | |
| 443 | unfolding Liminf_Sup top_ereal_def[symmetric] Sup_eq_top_iff | |
| 43920 | 444 | by (auto elim!: allE[where x="ereal m"] simp: top_ereal_def) | 
| 49664 | 445 | ultimately | 
| 446 | have "eventually (%x. f x : S) net" | |
| 447 | apply (subst eventually_mono) | |
| 448 | apply auto | |
| 449 | done | |
| 450 | } | |
| 451 | then show "(f ---> \<infinity>) net" unfolding tendsto_def by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 452 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 453 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 454 | lemma Limsup_MInfty: | 
| 43920 | 455 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 456 | assumes "\<not> trivial_limit net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 457 | shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>" | 
| 43920 | 458 | using assms ereal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"] | 
| 459 | ereal_Liminf_uminus[of _ f] by (auto simp: ereal_uminus_eq_reorder) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 460 | |
| 43920 | 461 | lemma ereal_Liminf_eq_Limsup: | 
| 462 | fixes f :: "'a \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 463 | assumes ntriv: "\<not> trivial_limit net" | 
| 49664 | 464 | and lim: "Liminf net f = f0" "Limsup net f = f0" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 465 | shows "(f ---> f0) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 466 | proof (cases f0) | 
| 49664 | 467 | case PInf | 
| 468 | then show ?thesis using Liminf_PInfty[OF ntriv] lim by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 469 | next | 
| 49664 | 470 | case MInf | 
| 471 | then show ?thesis using Limsup_MInfty[OF ntriv] lim by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 472 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 473 | case (real r) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 474 | show "(f ---> f0) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 475 | proof (rule topological_tendstoI) | 
| 49664 | 476 | fix S | 
| 477 | assume "open S""f0 \<in> S" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 478 |     then obtain a b where "a < Liminf net f" "Limsup net f < b" "{a<..<b} \<subseteq> S"
 | 
| 43920 | 479 | using ereal_open_cont_interval2[of S f0] real lim by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 480 |     then have "eventually (\<lambda>x. f x \<in> {a<..<b}) net"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 481 | unfolding Liminf_Sup Limsup_Inf less_Sup_iff Inf_less_iff | 
| 44142 | 482 | by (auto intro!: eventually_conj) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 483 |     with `{a<..<b} \<subseteq> S` show "eventually (%x. f x : S) net"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 484 | by (rule_tac eventually_mono) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 485 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 486 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 487 | |
| 43920 | 488 | lemma ereal_Liminf_eq_Limsup_iff: | 
| 489 | fixes f :: "'a \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 490 | assumes "\<not> trivial_limit net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 491 | shows "(f ---> f0) net \<longleftrightarrow> Liminf net f = f0 \<and> Limsup net f = f0" | 
| 43920 | 492 | by (metis assms ereal_Liminf_eq_Limsup lim_imp_Liminf lim_imp_Limsup) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 493 | |
| 50104 | 494 | lemma convergent_ereal: | 
| 495 | fixes X :: "nat \<Rightarrow> ereal" | |
| 496 | shows "convergent X \<longleftrightarrow> limsup X = liminf X" | |
| 497 | using ereal_Liminf_eq_Limsup_iff[of sequentially] | |
| 498 | by (auto simp: convergent_def) | |
| 499 | ||
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 500 | lemma limsup_INFI_SUPR: | 
| 43920 | 501 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 502 |   shows "limsup f = (INF n. SUP m:{n..}. f m)"
 | 
| 43920 | 503 | using ereal_Limsup_uminus[of sequentially "\<lambda>x. - f x"] | 
| 504 | by (simp add: liminf_SUPR_INFI ereal_INFI_uminus ereal_SUPR_uminus) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 505 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 506 | lemma liminf_PInfty: | 
| 43920 | 507 | fixes X :: "nat => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 508 | shows "X ----> \<infinity> <-> liminf X = \<infinity>" | 
| 49664 | 509 | by (metis Liminf_PInfty trivial_limit_sequentially) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 510 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 511 | lemma limsup_MInfty: | 
| 43920 | 512 | fixes X :: "nat => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 513 | shows "X ----> (-\<infinity>) <-> limsup X = (-\<infinity>)" | 
| 49664 | 514 | by (metis Limsup_MInfty trivial_limit_sequentially) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 515 | |
| 43920 | 516 | lemma ereal_lim_mono: | 
| 517 | fixes X Y :: "nat => ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 518 | assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n" | 
| 49664 | 519 | and "X ----> x" "Y ----> y" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 520 | shows "x <= y" | 
| 43920 | 521 | by (metis ereal_Liminf_eq_Limsup_iff[OF trivial_limit_sequentially] assms liminf_mono) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 522 | |
| 43920 | 523 | lemma incseq_le_ereal: | 
| 524 | fixes X :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 525 | assumes inc: "incseq X" and lim: "X ----> L" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 526 | shows "X N \<le> L" | 
| 49664 | 527 | using inc by (intro ereal_lim_mono[of N, OF _ tendsto_const lim]) (simp add: incseq_def) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 528 | |
| 49664 | 529 | lemma decseq_ge_ereal: | 
| 530 | assumes dec: "decseq X" | |
| 531 | and lim: "X ----> (L::ereal)" | |
| 532 | shows "X N >= L" | |
| 533 | using dec by (intro ereal_lim_mono[of N, OF _ lim tendsto_const]) (simp add: decseq_def) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 534 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 535 | lemma liminf_bounded_open: | 
| 43920 | 536 | fixes x :: "nat \<Rightarrow> ereal" | 
| 49664 | 537 | shows "x0 \<le> liminf x \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> x0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. x n \<in> S))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 538 | (is "_ \<longleftrightarrow> ?P x0") | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 539 | proof | 
| 49664 | 540 | assume "?P x0" | 
| 541 | then show "x0 \<le> liminf x" | |
| 43920 | 542 | unfolding ereal_Liminf_Sup_monoset eventually_sequentially | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 543 | by (intro complete_lattice_class.Sup_upper) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 544 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 545 | assume "x0 \<le> liminf x" | 
| 49664 | 546 |   { fix S :: "ereal set"
 | 
| 547 | assume om: "open S & mono_set S & x0:S" | |
| 548 |     { assume "S = UNIV" then have "EX N. (ALL n>=N. x n : S)" by auto }
 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 549 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 550 |     { assume "~(S=UNIV)"
 | 
| 43920 | 551 |       then obtain B where B_def: "S = {B<..}" using om ereal_open_mono_set by auto
 | 
| 49664 | 552 | then have "B<x0" using om by auto | 
| 553 | then have "EX N. ALL n>=N. x n : S" | |
| 554 | unfolding B_def using `x0 \<le> liminf x` liminf_bounded_iff by auto | |
| 555 | } | |
| 556 | ultimately have "EX N. (ALL n>=N. x n : S)" by auto | |
| 557 | } | |
| 558 | then show "?P x0" by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 559 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 560 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 561 | lemma limsup_subseq_mono: | 
| 43920 | 562 | fixes X :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 563 | assumes "subseq r" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 564 | shows "limsup (X \<circ> r) \<le> limsup X" | 
| 49664 | 565 | proof - | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 566 | have "(\<lambda>n. - X n) \<circ> r = (\<lambda>n. - (X \<circ> r) n)" by (simp add: fun_eq_iff) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 567 | then have "- limsup X \<le> - limsup (X \<circ> r)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 568 | using liminf_subseq_mono[of r "(%n. - X n)"] | 
| 43920 | 569 | ereal_Liminf_uminus[of sequentially X] | 
| 570 | ereal_Liminf_uminus[of sequentially "X o r"] assms by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 571 | then show ?thesis by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 572 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 573 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 574 | lemma bounded_abs: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 575 | assumes "(a::real)<=x" "x<=b" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 576 | shows "abs x <= max (abs a) (abs b)" | 
| 49664 | 577 | by (metis abs_less_iff assms leI le_max_iff_disj | 
| 578 | less_eq_real_def less_le_not_le less_minus_iff minus_minus) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 579 | |
| 49664 | 580 | lemma bounded_increasing_convergent2: | 
| 581 | fixes f::"nat => real" | |
| 582 | assumes "ALL n. f n <= B" "ALL n m. n>=m --> f n >= f m" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 583 | shows "EX l. (f ---> l) sequentially" | 
| 49664 | 584 | proof - | 
| 585 | def N == "max (abs (f 0)) (abs B)" | |
| 586 |   { fix n
 | |
| 587 | have "abs (f n) <= N" | |
| 588 | unfolding N_def | |
| 589 | apply (subst bounded_abs) | |
| 590 | using assms apply auto | |
| 591 | done } | |
| 592 |   then have "bounded {f n| n::nat. True}"
 | |
| 593 | unfolding bounded_real by auto | |
| 594 | then show ?thesis | |
| 595 | apply (rule Topology_Euclidean_Space.bounded_increasing_convergent) | |
| 596 | using assms apply auto | |
| 597 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 598 | qed | 
| 49664 | 599 | |
| 600 | lemma lim_ereal_increasing: | |
| 601 | assumes "\<And>n m. n >= m \<Longrightarrow> f n >= f m" | |
| 43920 | 602 | obtains l where "f ----> (l::ereal)" | 
| 49664 | 603 | proof (cases "f = (\<lambda>x. - \<infinity>)") | 
| 604 | case True | |
| 605 | then show thesis | |
| 606 | using tendsto_const[of "- \<infinity>" sequentially] by (intro that[of "-\<infinity>"]) auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 607 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 608 | case False | 
| 49664 | 609 | then obtain N where N_def: "f N > (-\<infinity>)" by (auto simp: fun_eq_iff) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 610 | have "ALL n>=N. f n >= f N" using assms by auto | 
| 49664 | 611 | then have minf: "ALL n>=N. f n > (-\<infinity>)" using N_def by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 612 | def Y == "(%n. (if n>=N then f n else f N))" | 
| 49664 | 613 | then have incy: "!!n m. n>=m ==> Y n >= Y m" using assms by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 614 | from minf have minfy: "ALL n. Y n ~= (-\<infinity>)" using Y_def by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 615 | show thesis | 
| 49664 | 616 | proof (cases "EX B. ALL n. f n < ereal B") | 
| 617 | case False | |
| 618 | then show thesis | |
| 619 | apply - | |
| 620 | apply (rule that[of \<infinity>]) | |
| 621 | unfolding Lim_PInfty not_ex not_all | |
| 622 | apply safe | |
| 623 | apply (erule_tac x=B in allE, safe) | |
| 624 | apply (rule_tac x=x in exI, safe) | |
| 625 | apply (rule order_trans[OF _ assms[rule_format]]) | |
| 626 | apply auto | |
| 627 | done | |
| 628 | next | |
| 629 | case True | |
| 630 | then guess B .. | |
| 631 | then have "ALL n. Y n < ereal B" using Y_def by auto | |
| 632 | note B = this[rule_format] | |
| 633 |     { fix n
 | |
| 634 | have "Y n < \<infinity>" | |
| 635 | using B[of n] | |
| 636 | apply (subst less_le_trans) | |
| 637 | apply auto | |
| 638 | done | |
| 639 | then have "Y n ~= \<infinity> & Y n ~= (-\<infinity>)" using minfy by auto | |
| 640 | } | |
| 641 | then have *: "ALL n. \<bar>Y n\<bar> \<noteq> \<infinity>" by auto | |
| 642 |     { fix n
 | |
| 643 | have "real (Y n) < B" | |
| 644 | proof - | |
| 645 | case goal1 | |
| 646 | then show ?case | |
| 647 | using B[of n] apply-apply(subst(asm) ereal_real'[THEN sym]) defer defer | |
| 648 | unfolding ereal_less using * by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 649 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 650 | } | 
| 49664 | 651 | then have B': "ALL n. (real (Y n) <= B)" using less_imp_le by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 652 | have "EX l. (%n. real (Y n)) ----> l" | 
| 49664 | 653 | apply (rule bounded_increasing_convergent2) | 
| 654 | proof safe | |
| 655 | show "\<And>n. real (Y n) <= B" using B' by auto | |
| 656 | fix n m :: nat | |
| 657 | assume "n<=m" | |
| 658 | then have "ereal (real (Y n)) <= ereal (real (Y m))" | |
| 43920 | 659 | using incy[rule_format,of n m] apply(subst ereal_real)+ | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 660 | using *[rule_format, of n] *[rule_format, of m] by auto | 
| 49664 | 661 | then show "real (Y n) <= real (Y m)" by auto | 
| 662 | qed | |
| 663 | then guess l .. note l=this | |
| 664 | have "Y ----> ereal l" | |
| 665 | using l | |
| 666 | apply - | |
| 667 | apply (subst(asm) lim_ereal[THEN sym]) | |
| 668 | unfolding ereal_real | |
| 669 | using * apply auto | |
| 670 | done | |
| 671 | then show thesis | |
| 672 | apply - | |
| 673 | apply (rule that[of "ereal l"]) | |
| 674 | apply (subst tail_same_limit[of Y _ N]) | |
| 675 | using Y_def apply auto | |
| 676 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 677 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 678 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 679 | |
| 49664 | 680 | lemma lim_ereal_decreasing: | 
| 681 | assumes "\<And>n m. n >= m \<Longrightarrow> f n <= f m" | |
| 43920 | 682 | obtains l where "f ----> (l::ereal)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 683 | proof - | 
| 43920 | 684 | from lim_ereal_increasing[of "\<lambda>x. - f x"] assms | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 685 | obtain l where "(\<lambda>x. - f x) ----> l" by auto | 
| 43920 | 686 | from ereal_lim_mult[OF this, of "- 1"] show thesis | 
| 687 | by (intro that[of "-l"]) (simp add: ereal_uminus_eq_reorder) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 688 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 689 | |
| 43920 | 690 | lemma compact_ereal: | 
| 691 | fixes X :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 692 | shows "\<exists>l r. subseq r \<and> (X \<circ> r) ----> l" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 693 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 694 | obtain r where "subseq r" and mono: "monoseq (X \<circ> r)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 695 | using seq_monosub[of X] unfolding comp_def by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 696 | then have "(\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) m \<le> (X \<circ> r) n) \<or> (\<forall>n m. m \<le> n \<longrightarrow> (X \<circ> r) n \<le> (X \<circ> r) m)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 697 | by (auto simp add: monoseq_def) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 698 | then obtain l where "(X\<circ>r) ----> l" | 
| 43920 | 699 | using lim_ereal_increasing[of "X \<circ> r"] lim_ereal_decreasing[of "X \<circ> r"] by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 700 | then show ?thesis using `subseq r` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 701 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 702 | |
| 43920 | 703 | lemma ereal_Sup_lim: | 
| 704 | assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 705 | shows "a \<le> Sup s" | 
| 49664 | 706 | by (metis Lim_bounded_ereal assms complete_lattice_class.Sup_upper) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 707 | |
| 43920 | 708 | lemma ereal_Inf_lim: | 
| 709 | assumes "\<And>n. b n \<in> s" "b ----> (a::ereal)" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 710 | shows "Inf s \<le> a" | 
| 49664 | 711 | by (metis Lim_bounded2_ereal assms complete_lattice_class.Inf_lower) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 712 | |
| 43920 | 713 | lemma SUP_Lim_ereal: | 
| 49664 | 714 | fixes X :: "nat \<Rightarrow> ereal" | 
| 715 | assumes "incseq X" "X ----> l" | |
| 716 | shows "(SUP n. X n) = l" | |
| 43920 | 717 | proof (rule ereal_SUPI) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 718 | fix n from assms show "X n \<le> l" | 
| 43920 | 719 | by (intro incseq_le_ereal) (simp add: incseq_def) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 720 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 721 | fix y assume "\<And>n. n \<in> UNIV \<Longrightarrow> X n \<le> y" | 
| 49664 | 722 |   with ereal_Sup_lim[OF _ `X ----> l`, of "{..y}"] show "l \<le> y" by auto
 | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 723 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 724 | |
| 43920 | 725 | lemma LIMSEQ_ereal_SUPR: | 
| 49664 | 726 | fixes X :: "nat \<Rightarrow> ereal" | 
| 727 | assumes "incseq X" | |
| 728 | shows "X ----> (SUP n. X n)" | |
| 43920 | 729 | proof (rule lim_ereal_increasing) | 
| 49664 | 730 | fix n m :: nat | 
| 731 | assume "m \<le> n" | |
| 732 | then show "X m \<le> X n" using `incseq X` by (simp add: incseq_def) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 733 | next | 
| 49664 | 734 | fix l | 
| 735 | assume "X ----> l" | |
| 43920 | 736 | with SUP_Lim_ereal[of X, OF assms this] show ?thesis by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 737 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 738 | |
| 43920 | 739 | lemma INF_Lim_ereal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::ereal)" | 
| 740 | using SUP_Lim_ereal[of "\<lambda>i. - X i" "- l"] | |
| 741 | by (simp add: ereal_SUPR_uminus ereal_lim_uminus) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 742 | |
| 43920 | 743 | lemma LIMSEQ_ereal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: ereal)" | 
| 744 | using LIMSEQ_ereal_SUPR[of "\<lambda>i. - X i"] | |
| 745 | by (simp add: ereal_SUPR_uminus ereal_lim_uminus) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 746 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 747 | lemma SUP_eq_LIMSEQ: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 748 | assumes "mono f" | 
| 43920 | 749 | shows "(SUP n. ereal (f n)) = ereal x \<longleftrightarrow> f ----> x" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 750 | proof | 
| 43920 | 751 | have inc: "incseq (\<lambda>i. ereal (f i))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 752 | using `mono f` unfolding mono_def incseq_def by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 753 |   { assume "f ----> x"
 | 
| 49664 | 754 | then have "(\<lambda>i. ereal (f i)) ----> ereal x" by auto | 
| 755 | from SUP_Lim_ereal[OF inc this] | |
| 756 | show "(SUP n. ereal (f n)) = ereal x" . } | |
| 43920 | 757 |   { assume "(SUP n. ereal (f n)) = ereal x"
 | 
| 758 | with LIMSEQ_ereal_SUPR[OF inc] | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 759 | show "f ----> x" by auto } | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 760 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 761 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 762 | lemma Liminf_within: | 
| 43920 | 763 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 764 |   shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
 | 
| 49664 | 765 | proof - | 
| 766 |   let ?l="(SUP e:{0<..}. INF y:(S Int ball x e - {x}). f y)"
 | |
| 767 |   { fix T
 | |
| 768 | assume T_def: "open T & mono_set T & ?l:T" | |
| 769 | have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T" | |
| 770 | proof - | |
| 771 |       { assume "T=UNIV" then have ?thesis by (simp add: gt_ex) }
 | |
| 772 | moreover | |
| 773 |       { assume "~(T=UNIV)"
 | |
| 774 |         then obtain B where "T={B<..}" using T_def ereal_open_mono_set[of T] by auto
 | |
| 775 | then have "B<?l" using T_def by auto | |
| 776 |         then obtain d where d_def: "0<d & B<(INF y:(S Int ball x d - {x}). f y)"
 | |
| 777 | unfolding less_SUP_iff by auto | |
| 778 |         { fix y assume "y:S & 0 < dist y x & dist y x < d"
 | |
| 779 |           then have "y:(S Int ball x d - {x})" unfolding ball_def by (auto simp add: dist_commute)
 | |
| 780 |           then have "f y:T" using d_def INF_lower[of y "S Int ball x d - {x}" f] `T={B<..}` by auto
 | |
| 781 | } then have ?thesis apply(rule_tac x="d" in exI) using d_def by auto | |
| 782 | } | |
| 783 | ultimately show ?thesis by auto | |
| 784 | qed | |
| 785 | } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 786 | moreover | 
| 49664 | 787 |   { fix z
 | 
| 788 | assume a: "ALL T. open T --> mono_set T --> z : T --> | |
| 789 | (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)" | |
| 790 |     { fix B
 | |
| 791 | assume "B<z" | |
| 792 | then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> B < f y)" | |
| 793 |          using a[rule_format, of "{B<..}"] mono_greaterThan by auto
 | |
| 794 |       { fix y
 | |
| 795 |         assume "y:(S Int ball x d - {x})"
 | |
| 796 | then have "y:S & 0 < dist y x & dist y x < d" | |
| 797 | unfolding ball_def | |
| 798 | apply (simp add: dist_commute) | |
| 799 | apply (metis dist_eq_0_iff less_le zero_le_dist) | |
| 800 | done | |
| 801 | then have "B <= f y" using d_def by auto | |
| 802 | } | |
| 803 |       then have "B <= INFI (S Int ball x d - {x}) f"
 | |
| 804 | apply (subst INF_greatest) | |
| 805 | apply auto | |
| 806 | done | |
| 807 | also have "...<=?l" | |
| 808 | apply (subst SUP_upper) | |
| 809 | using d_def apply auto | |
| 810 | done | |
| 811 | finally have "B<=?l" by auto | |
| 812 | } | |
| 813 | then have "z <= ?l" using ereal_le_ereal[of z "?l"] by auto | |
| 814 | } | |
| 815 | ultimately show ?thesis | |
| 816 | unfolding ereal_Liminf_Sup_monoset eventually_within | |
| 817 |     apply (subst ereal_SupI[of _ "(SUP e:{0<..}. INFI (S Int ball x e - {x}) f)"])
 | |
| 818 | apply auto | |
| 819 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 820 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 821 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 822 | lemma Limsup_within: | 
| 43920 | 823 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 824 |   shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
 | 
| 49664 | 825 | proof - | 
| 826 |   let ?l = "(INF e:{0<..}. SUP y:(S Int ball x e - {x}). f y)"
 | |
| 827 |   { fix T
 | |
| 828 | assume T_def: "open T & mono_set (uminus ` T) & ?l:T" | |
| 829 | have "EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T" | |
| 830 | proof - | |
| 831 |       { assume "T = UNIV"
 | |
| 832 | then have ?thesis by (simp add: gt_ex) } | |
| 833 | moreover | |
| 834 |       { assume "T \<noteq> UNIV"
 | |
| 835 | then have "~(uminus ` T = UNIV)" | |
| 836 | by (metis Int_UNIV_right Int_absorb1 image_mono ereal_minus_minus_image subset_UNIV) | |
| 837 |         then have "uminus ` T = {Inf (uminus ` T)<..}"
 | |
| 838 | using T_def ereal_open_mono_set[of "uminus ` T"] ereal_open_uminus[of T] by auto | |
| 839 |         then obtain B where "T={..<B}"
 | |
| 840 | unfolding ereal_Inf_uminus_image_eq ereal_uminus_lessThan[symmetric] | |
| 841 | unfolding inj_image_eq_iff[OF ereal_inj_on_uminus] by simp | |
| 842 | then have "?l<B" using T_def by auto | |
| 843 |         then obtain d where d_def: "0<d & (SUP y:(S Int ball x d - {x}). f y)<B"
 | |
| 844 | unfolding INF_less_iff by auto | |
| 845 |         { fix y
 | |
| 846 | assume "y:S & 0 < dist y x & dist y x < d" | |
| 847 |           then have "y:(S Int ball x d - {x})"
 | |
| 848 | unfolding ball_def by (auto simp add: dist_commute) | |
| 849 | then have "f y:T" | |
| 850 |             using d_def SUP_upper[of y "S Int ball x d - {x}" f] `T={..<B}` by auto
 | |
| 851 | } | |
| 852 | then have ?thesis | |
| 853 | apply (rule_tac x="d" in exI) | |
| 854 | using d_def apply auto | |
| 855 | done | |
| 856 | } | |
| 857 | ultimately show ?thesis by auto | |
| 858 | qed | |
| 859 | } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 860 | moreover | 
| 49664 | 861 |   { fix z
 | 
| 862 | assume a: "ALL T. open T --> mono_set (uminus ` T) --> z : T --> | |
| 863 | (EX d>0. ALL y:S. 0 < dist y x & dist y x < d --> f y : T)" | |
| 864 |     { fix B
 | |
| 865 | assume "z<B" | |
| 866 | then obtain d where d_def: "d>0 & (ALL y:S. 0 < dist y x & dist y x < d --> f y<B)" | |
| 867 |          using a[rule_format, of "{..<B}"] by auto
 | |
| 868 |       { fix y
 | |
| 869 |         assume "y:(S Int ball x d - {x})"
 | |
| 870 | then have "y:S & 0 < dist y x & dist y x < d" | |
| 871 | unfolding ball_def | |
| 872 | apply (simp add: dist_commute) | |
| 873 | apply (metis dist_eq_0_iff less_le zero_le_dist) | |
| 874 | done | |
| 875 | then have "f y <= B" using d_def by auto | |
| 876 | } | |
| 877 |       then have "SUPR (S Int ball x d - {x}) f <= B"
 | |
| 878 | apply (subst SUP_least) | |
| 879 | apply auto | |
| 880 | done | |
| 881 | moreover | |
| 882 |       have "?l<=SUPR (S Int ball x d - {x}) f"
 | |
| 883 | apply (subst INF_lower) | |
| 884 | using d_def apply auto | |
| 885 | done | |
| 886 | ultimately have "?l<=B" by auto | |
| 887 | } then have "?l <= z" using ereal_ge_ereal[of z "?l"] by auto | |
| 888 | } | |
| 889 | ultimately show ?thesis | |
| 890 | unfolding ereal_Limsup_Inf_monoset eventually_within | |
| 891 | apply (subst ereal_InfI) | |
| 892 | apply auto | |
| 893 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 894 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 895 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 896 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 897 | lemma Liminf_within_UNIV: | 
| 43920 | 898 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 899 | shows "Liminf (at x) f = Liminf (at x within UNIV) f" | 
| 45031 | 900 | by simp (* TODO: delete *) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 901 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 902 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 903 | lemma Liminf_at: | 
| 43920 | 904 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 905 |   shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
 | 
| 45031 | 906 | using Liminf_within[of x UNIV f] by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 907 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 908 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 909 | lemma Limsup_within_UNIV: | 
| 43920 | 910 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 911 | shows "Limsup (at x) f = Limsup (at x within UNIV) f" | 
| 45031 | 912 | by simp (* TODO: delete *) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 913 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 914 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 915 | lemma Limsup_at: | 
| 43920 | 916 | fixes f :: "'a::metric_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 917 |   shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
 | 
| 45031 | 918 | using Limsup_within[of x UNIV f] by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 919 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 920 | lemma Lim_within_constant: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 921 | assumes "ALL y:S. f y = C" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 922 | shows "(f ---> C) (at x within S)" | 
| 45032 | 923 | unfolding tendsto_def Limits.eventually_within eventually_at_topological | 
| 924 | using assms by simp (metis open_UNIV UNIV_I) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 925 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 926 | lemma Liminf_within_constant: | 
| 45032 | 927 | fixes f :: "'a::topological_space \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 928 | assumes "ALL y:S. f y = C" | 
| 49664 | 929 | and "~trivial_limit (at x within S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 930 | shows "Liminf (at x within S) f = C" | 
| 49664 | 931 | by (metis Lim_within_constant assms lim_imp_Liminf) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 932 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 933 | lemma Limsup_within_constant: | 
| 45032 | 934 | fixes f :: "'a::topological_space \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 935 | assumes "ALL y:S. f y = C" | 
| 49664 | 936 | and "~trivial_limit (at x within S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 937 | shows "Limsup (at x within S) f = C" | 
| 49664 | 938 | by (metis Lim_within_constant assms lim_imp_Limsup) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 939 | |
| 49664 | 940 | lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
 | 
| 941 | unfolding islimpt_def by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 942 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 943 | |
| 49664 | 944 | lemma islimpt_in_closure: "(x islimpt S) = (x:closure(S-{x}))"
 | 
| 945 | unfolding closure_def using islimpt_punctured by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 946 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 947 | |
| 49664 | 948 | lemma not_trivial_limit_within: "~trivial_limit (at x within S) = (x:closure(S-{x}))"
 | 
| 949 | using islimpt_in_closure by (metis trivial_limit_within) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 950 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 951 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 952 | lemma not_trivial_limit_within_ball: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 953 |   "(~trivial_limit (at x within S)) = (ALL e>0. S Int ball x e - {x} ~= {})"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 954 | (is "?lhs = ?rhs") | 
| 49664 | 955 | proof - | 
| 956 |   { assume "?lhs"
 | |
| 957 |     { fix e :: real
 | |
| 958 | assume "e>0" | |
| 959 |       then obtain y where "y:(S-{x}) & dist y x < e"
 | |
| 960 |         using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
 | |
| 961 | by auto | |
| 962 |       then have "y : (S Int ball x e - {x})"
 | |
| 963 | unfolding ball_def by (simp add: dist_commute) | |
| 964 |       then have "S Int ball x e - {x} ~= {}" by blast
 | |
| 965 | } then have "?rhs" by auto | |
| 966 | } | |
| 967 | moreover | |
| 968 |   { assume "?rhs"
 | |
| 969 |     { fix e :: real
 | |
| 970 | assume "e>0" | |
| 971 |       then obtain y where "y : (S Int ball x e - {x})" using `?rhs` by blast
 | |
| 972 |       then have "y:(S-{x}) & dist y x < e"
 | |
| 973 | unfolding ball_def by (simp add: dist_commute) | |
| 974 |       then have "EX y:(S-{x}). dist y x < e" by auto
 | |
| 975 | } | |
| 976 | then have "?lhs" | |
| 977 |       using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto
 | |
| 978 | } | |
| 979 | ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 980 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 981 | |
| 43920 | 982 | lemma liminf_ereal_cminus: | 
| 49664 | 983 | fixes f :: "nat \<Rightarrow> ereal" | 
| 984 | assumes "c \<noteq> -\<infinity>" | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 985 | shows "liminf (\<lambda>x. c - f x) = c - limsup f" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 986 | proof (cases c) | 
| 49664 | 987 | case PInf | 
| 988 | then show ?thesis by (simp add: Liminf_const) | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 989 | next | 
| 49664 | 990 | case (real r) | 
| 991 | then show ?thesis | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 992 | unfolding liminf_SUPR_INFI limsup_INFI_SUPR | 
| 43920 | 993 | apply (subst INFI_ereal_cminus) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 994 | apply auto | 
| 43920 | 995 | apply (subst SUPR_ereal_cminus) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 996 | apply auto | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 997 | done | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 998 | qed (insert `c \<noteq> -\<infinity>`, simp) | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 999 | |
| 49664 | 1000 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1001 | subsubsection {* Continuity *}
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1002 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1003 | lemma continuous_imp_tendsto: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1004 | assumes "continuous (at x0) f" | 
| 49664 | 1005 | and "x ----> x0" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1006 | shows "(f o x) ----> (f x0)" | 
| 49664 | 1007 | proof - | 
| 1008 |   { fix S
 | |
| 1009 | assume "open S & (f x0):S" | |
| 1010 | then obtain T where T_def: "open T & x0 : T & (ALL x:T. f x : S)" | |
| 1011 | using assms continuous_at_open by metis | |
| 1012 | then have "(EX N. ALL n>=N. x n : T)" | |
| 1013 | using assms tendsto_explicit T_def by auto | |
| 1014 | then have "(EX N. ALL n>=N. f(x n) : S)" using T_def by auto | |
| 1015 | } | |
| 1016 | then show ?thesis using tendsto_explicit[of "f o x" "f x0"] by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1017 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1018 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1019 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1020 | lemma continuous_at_sequentially2: | 
| 49664 | 1021 | fixes f :: "'a::metric_space => 'b:: topological_space" | 
| 1022 | shows "continuous (at x0) f <-> (ALL x. (x ----> x0) --> (f o x) ----> (f x0))" | |
| 1023 | proof - | |
| 1024 |   { assume "~(continuous (at x0) f)"
 | |
| 1025 | then obtain T where | |
| 1026 | T_def: "open T & f x0 : T & (ALL S. (open S & x0 : S) --> (EX x':S. f x' ~: T))" | |
| 1027 | using continuous_at_open[of x0 f] by metis | |
| 1028 |     def X == "{x'. f x' ~: T}"
 | |
| 1029 | then have "x0 islimpt X" | |
| 1030 | unfolding islimpt_def using T_def by auto | |
| 1031 | then obtain x where x_def: "(ALL n. x n : X) & x ----> x0" | |
| 1032 | using islimpt_sequential[of x0 X] by auto | |
| 1033 | then have "~(f o x) ----> (f x0)" | |
| 1034 | unfolding tendsto_explicit using X_def T_def by auto | |
| 1035 | then have "EX x. x ----> x0 & (~(f o x) ----> (f x0))" using x_def by auto | |
| 1036 | } | |
| 1037 | then show ?thesis using continuous_imp_tendsto by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1038 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1039 | |
| 43920 | 1040 | lemma continuous_at_of_ereal: | 
| 1041 | fixes x0 :: ereal | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1042 | assumes "\<bar>x0\<bar> \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1043 | shows "continuous (at x0) real" | 
| 49664 | 1044 | proof - | 
| 1045 |   { fix T
 | |
| 1046 | assume T_def: "open T & real x0 : T" | |
| 1047 | def S == "ereal ` T" | |
| 1048 | then have "ereal (real x0) : S" using T_def by auto | |
| 1049 | then have "x0 : S" using assms ereal_real by auto | |
| 1050 | moreover have "open S" using open_ereal S_def T_def by auto | |
| 1051 | moreover have "ALL y:S. real y : T" using S_def T_def by auto | |
| 1052 | ultimately have "EX S. x0 : S & open S & (ALL y:S. real y : T)" by auto | |
| 1053 | } | |
| 1054 | then show ?thesis unfolding continuous_at_open by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1055 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1056 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1057 | |
| 43920 | 1058 | lemma continuous_at_iff_ereal: | 
| 49664 | 1059 | fixes f :: "'a::t2_space => real" | 
| 1060 | shows "continuous (at x0) f <-> continuous (at x0) (ereal o f)" | |
| 1061 | proof - | |
| 1062 |   { assume "continuous (at x0) f"
 | |
| 1063 | then have "continuous (at x0) (ereal o f)" | |
| 1064 | using continuous_at_ereal continuous_at_compose[of x0 f ereal] by auto | |
| 1065 | } | |
| 1066 | moreover | |
| 1067 |   { assume "continuous (at x0) (ereal o f)"
 | |
| 1068 | then have "continuous (at x0) (real o (ereal o f))" | |
| 1069 | using continuous_at_of_ereal by (intro continuous_at_compose[of x0 "ereal o f"]) auto | |
| 1070 | moreover have "real o (ereal o f) = f" using real_ereal_id by (simp add: o_assoc) | |
| 1071 | ultimately have "continuous (at x0) f" by auto | |
| 1072 | } ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1073 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1074 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1075 | |
| 43920 | 1076 | lemma continuous_on_iff_ereal: | 
| 49664 | 1077 | fixes f :: "'a::t2_space => real" | 
| 1078 | fixes A assumes "open A" | |
| 1079 | shows "continuous_on A f <-> continuous_on A (ereal o f)" | |
| 1080 | using continuous_at_iff_ereal assms by (auto simp add: continuous_on_eq_continuous_at) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1081 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1082 | |
| 43923 | 1083 | lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>::ereal)}) real"
 | 
| 49664 | 1084 | using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1085 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1086 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1087 | lemma continuous_on_iff_real: | 
| 43920 | 1088 | fixes f :: "'a::t2_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1089 | assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1090 | shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)" | 
| 49664 | 1091 | proof - | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1092 |   have "f ` A <= UNIV-{\<infinity>,(-\<infinity>)}" using assms by force
 | 
| 49664 | 1093 | then have *: "continuous_on (f ` A) real" | 
| 1094 | using continuous_on_real by (simp add: continuous_on_subset) | |
| 1095 | have **: "continuous_on ((real o f) ` A) ereal" | |
| 1096 | using continuous_on_ereal continuous_on_subset[of "UNIV" "ereal" "(real o f) ` A"] by blast | |
| 1097 |   { assume "continuous_on A f"
 | |
| 1098 | then have "continuous_on A (real o f)" | |
| 1099 | apply (subst continuous_on_compose) | |
| 1100 | using * apply auto | |
| 1101 | done | |
| 1102 | } | |
| 1103 | moreover | |
| 1104 |   { assume "continuous_on A (real o f)"
 | |
| 1105 | then have "continuous_on A (ereal o (real o f))" | |
| 1106 | apply (subst continuous_on_compose) | |
| 1107 | using ** apply auto | |
| 1108 | done | |
| 1109 | then have "continuous_on A f" | |
| 1110 | apply (subst continuous_on_eq[of A "ereal o (real o f)" f]) | |
| 1111 | using assms ereal_real apply auto | |
| 1112 | done | |
| 1113 | } | |
| 1114 | ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1115 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1116 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1117 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1118 | lemma continuous_at_const: | 
| 43920 | 1119 | fixes f :: "'a::t2_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1120 | assumes "ALL x. (f x = C)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1121 | shows "ALL x. continuous (at x) f" | 
| 49664 | 1122 | unfolding continuous_at_open using assms t1_space by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1123 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1124 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1125 | lemma closure_contains_Inf: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1126 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1127 |   assumes "S ~= {}" "EX B. ALL x:S. B<=x"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1128 | shows "Inf S : closure S" | 
| 49664 | 1129 | proof - | 
| 1130 | have *: "ALL x:S. Inf S <= x" | |
| 1131 | using Inf_lower_EX[of _ S] assms by metis | |
| 1132 |   { fix e
 | |
| 1133 | assume "e>(0 :: real)" | |
| 1134 |     then obtain x where x_def: "x:S & x < Inf S + e" using Inf_close `S ~= {}` by auto
 | |
| 1135 | moreover then have "x > Inf S - e" using * by auto | |
| 1136 | ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff) | |
| 1137 | then have "EX x:S. abs (x - Inf S) < e" using x_def by auto | |
| 1138 | } | |
| 1139 | then show ?thesis | |
| 1140 | apply (subst closure_approachable) | |
| 1141 | unfolding dist_norm apply auto | |
| 1142 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1143 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1144 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1145 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1146 | lemma closed_contains_Inf: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1147 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1148 |   assumes "S ~= {}" "EX B. ALL x:S. B<=x"
 | 
| 49664 | 1149 | and "closed S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1150 | shows "Inf S : S" | 
| 49664 | 1151 | by (metis closure_contains_Inf closure_closed assms) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1152 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1153 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1154 | lemma mono_closed_real: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1155 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1156 | assumes mono: "ALL y z. y:S & y<=z --> z:S" | 
| 49664 | 1157 | and "closed S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1158 |   shows "S = {} | S = UNIV | (EX a. S = {a ..})"
 | 
| 49664 | 1159 | proof - | 
| 1160 |   { assume "S ~= {}"
 | |
| 1161 |     { assume ex: "EX B. ALL x:S. B<=x"
 | |
| 1162 | then have *: "ALL x:S. Inf S <= x" using Inf_lower_EX[of _ S] ex by metis | |
| 1163 |       then have "Inf S : S" apply (subst closed_contains_Inf) using ex `S ~= {}` `closed S` by auto
 | |
| 1164 | then have "ALL x. (Inf S <= x <-> x:S)" using mono[rule_format, of "Inf S"] * by auto | |
| 1165 |       then have "S = {Inf S ..}" by auto
 | |
| 1166 |       then have "EX a. S = {a ..}" by auto
 | |
| 1167 | } | |
| 1168 | moreover | |
| 1169 |     { assume "~(EX B. ALL x:S. B<=x)"
 | |
| 1170 | then have nex: "ALL B. EX x:S. x<B" by (simp add: not_le) | |
| 1171 |       { fix y
 | |
| 1172 | obtain x where "x:S & x < y" using nex by auto | |
| 1173 | then have "y:S" using mono[rule_format, of x y] by auto | |
| 1174 | } then have "S = UNIV" by auto | |
| 1175 | } | |
| 1176 |     ultimately have "S = UNIV | (EX a. S = {a ..})" by blast
 | |
| 1177 | } then show ?thesis by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1178 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1179 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1180 | |
| 43920 | 1181 | lemma mono_closed_ereal: | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1182 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1183 | assumes mono: "ALL y z. y:S & y<=z --> z:S" | 
| 49664 | 1184 | and "closed S" | 
| 43920 | 1185 |   shows "EX a. S = {x. a <= ereal x}"
 | 
| 49664 | 1186 | proof - | 
| 1187 |   { assume "S = {}"
 | |
| 1188 | then have ?thesis apply(rule_tac x=PInfty in exI) by auto } | |
| 1189 | moreover | |
| 1190 |   { assume "S = UNIV"
 | |
| 1191 | then have ?thesis apply(rule_tac x="-\<infinity>" in exI) by auto } | |
| 1192 | moreover | |
| 1193 |   { assume "EX a. S = {a ..}"
 | |
| 1194 |     then obtain a where "S={a ..}" by auto
 | |
| 1195 | then have ?thesis apply(rule_tac x="ereal a" in exI) by auto | |
| 1196 | } | |
| 1197 | ultimately show ?thesis using mono_closed_real[of S] assms by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1198 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1199 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1200 | subsection {* Sums *}
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1201 | |
| 49664 | 1202 | lemma setsum_ereal[simp]: "(\<Sum>x\<in>A. ereal (f x)) = ereal (\<Sum>x\<in>A. f x)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1203 | proof cases | 
| 49664 | 1204 | assume "finite A" | 
| 1205 | then show ?thesis by induct auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1206 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1207 | |
| 43923 | 1208 | lemma setsum_Pinfty: | 
| 1209 | fixes f :: "'a \<Rightarrow> ereal" | |
| 1210 | shows "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<infinity>))" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1211 | proof safe | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1212 | assume *: "setsum f P = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1213 | show "finite P" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1214 | proof (rule ccontr) assume "infinite P" with * show False by auto qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1215 | show "\<exists>i\<in>P. f i = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1216 | proof (rule ccontr) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1217 | assume "\<not> ?thesis" then have "\<And>i. i \<in> P \<Longrightarrow> f i \<noteq> \<infinity>" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1218 | from `finite P` this have "setsum f P \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1219 | by induct auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1220 | with * show False by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1221 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1222 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1223 | fix i assume "finite P" "i \<in> P" "f i = \<infinity>" | 
| 49664 | 1224 | then show "setsum f P = \<infinity>" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1225 | proof induct | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1226 | case (insert x A) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1227 | show ?case using insert by (cases "x = i") auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1228 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1229 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1230 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1231 | lemma setsum_Inf: | 
| 43923 | 1232 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1233 | shows "\<bar>setsum f A\<bar> = \<infinity> \<longleftrightarrow> (finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>))" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1234 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1235 | assume *: "\<bar>setsum f A\<bar> = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1236 | have "finite A" by (rule ccontr) (insert *, auto) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1237 | moreover have "\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1238 | proof (rule ccontr) | 
| 43920 | 1239 | assume "\<not> ?thesis" then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1240 | from bchoice[OF this] guess r .. | 
| 44142 | 1241 | with * show False by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1242 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1243 | ultimately show "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1244 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1245 | assume "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1246 | then obtain i where "finite A" "i \<in> A" "\<bar>f i\<bar> = \<infinity>" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1247 | then show "\<bar>setsum f A\<bar> = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1248 | proof induct | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1249 | case (insert j A) then show ?case | 
| 43920 | 1250 | by (cases rule: ereal3_cases[of "f i" "f j" "setsum f A"]) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1251 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1252 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1253 | |
| 43920 | 1254 | lemma setsum_real_of_ereal: | 
| 43923 | 1255 | fixes f :: "'i \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1256 | assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1257 | shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1258 | proof - | 
| 43920 | 1259 | have "\<forall>x\<in>S. \<exists>r. f x = ereal r" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1260 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1261 | fix x assume "x \<in> S" | 
| 43920 | 1262 | from assms[OF this] show "\<exists>r. f x = ereal r" by (cases "f x") auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1263 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1264 | from bchoice[OF this] guess r .. | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1265 | then show ?thesis by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1266 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1267 | |
| 43920 | 1268 | lemma setsum_ereal_0: | 
| 1269 | fixes f :: "'a \<Rightarrow> ereal" assumes "finite A" "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1270 | shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1271 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1272 | assume *: "(\<Sum>x\<in>A. f x) = 0" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1273 | then have "(\<Sum>x\<in>A. f x) \<noteq> \<infinity>" by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1274 | then have "\<forall>i\<in>A. \<bar>f i\<bar> \<noteq> \<infinity>" using assms by (force simp: setsum_Pinfty) | 
| 43920 | 1275 | then have "\<forall>i\<in>A. \<exists>r. f i = ereal r" by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1276 | from bchoice[OF this] * assms show "\<forall>i\<in>A. f i = 0" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1277 | using setsum_nonneg_eq_0_iff[of A "\<lambda>i. real (f i)"] by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1278 | qed (rule setsum_0') | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1279 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1280 | |
| 43920 | 1281 | lemma setsum_ereal_right_distrib: | 
| 49664 | 1282 | fixes f :: "'a \<Rightarrow> ereal" | 
| 1283 | assumes "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1284 | shows "r * setsum f A = (\<Sum>n\<in>A. r * f n)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1285 | proof cases | 
| 49664 | 1286 | assume "finite A" | 
| 1287 | then show ?thesis using assms | |
| 43920 | 1288 | by induct (auto simp: ereal_right_distrib setsum_nonneg) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1289 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1290 | |
| 43920 | 1291 | lemma sums_ereal_positive: | 
| 49664 | 1292 | fixes f :: "nat \<Rightarrow> ereal" | 
| 1293 | assumes "\<And>i. 0 \<le> f i" | |
| 1294 | shows "f sums (SUP n. \<Sum>i<n. f i)" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1295 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1296 | have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)" | 
| 43920 | 1297 | using ereal_add_mono[OF _ assms] by (auto intro!: incseq_SucI) | 
| 1298 | from LIMSEQ_ereal_SUPR[OF this] | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1299 | show ?thesis unfolding sums_def by (simp add: atLeast0LessThan) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1300 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1301 | |
| 43920 | 1302 | lemma summable_ereal_pos: | 
| 49664 | 1303 | fixes f :: "nat \<Rightarrow> ereal" | 
| 1304 | assumes "\<And>i. 0 \<le> f i" | |
| 1305 | shows "summable f" | |
| 43920 | 1306 | using sums_ereal_positive[of f, OF assms] unfolding summable_def by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1307 | |
| 43920 | 1308 | lemma suminf_ereal_eq_SUPR: | 
| 49664 | 1309 | fixes f :: "nat \<Rightarrow> ereal" | 
| 1310 | assumes "\<And>i. 0 \<le> f i" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1311 | shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)" | 
| 43920 | 1312 | using sums_ereal_positive[of f, OF assms, THEN sums_unique] by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1313 | |
| 49664 | 1314 | lemma sums_ereal: "(\<lambda>x. ereal (f x)) sums ereal x \<longleftrightarrow> f sums x" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1315 | unfolding sums_def by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1316 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1317 | lemma suminf_bound: | 
| 43920 | 1318 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1319 | assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x" and pos: "\<And>n. 0 \<le> f n" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1320 | shows "suminf f \<le> x" | 
| 43920 | 1321 | proof (rule Lim_bounded_ereal) | 
| 1322 | have "summable f" using pos[THEN summable_ereal_pos] . | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1323 | then show "(\<lambda>N. \<Sum>n<N. f n) ----> suminf f" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1324 | by (auto dest!: summable_sums simp: sums_def atLeast0LessThan) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1325 |   show "\<forall>n\<ge>0. setsum f {..<n} \<le> x"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1326 | using assms by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1327 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1328 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1329 | lemma suminf_bound_add: | 
| 43920 | 1330 | fixes f :: "nat \<Rightarrow> ereal" | 
| 49664 | 1331 | assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x" | 
| 1332 | and pos: "\<And>n. 0 \<le> f n" | |
| 1333 | and "y \<noteq> -\<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1334 | shows "suminf f + y \<le> x" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1335 | proof (cases y) | 
| 49664 | 1336 | case (real r) | 
| 1337 | then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" | |
| 43920 | 1338 | using assms by (simp add: ereal_le_minus) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1339 | then have "(\<Sum> n. f n) \<le> x - y" using pos by (rule suminf_bound) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1340 | then show "(\<Sum> n. f n) + y \<le> x" | 
| 43920 | 1341 | using assms real by (simp add: ereal_le_minus) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1342 | qed (insert assms, auto) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1343 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1344 | lemma suminf_upper: | 
| 49664 | 1345 | fixes f :: "nat \<Rightarrow> ereal" | 
| 1346 | assumes "\<And>n. 0 \<le> f n" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1347 | shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)" | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44918diff
changeset | 1348 | unfolding suminf_ereal_eq_SUPR[OF assms] SUP_def | 
| 45031 | 1349 | by (auto intro: complete_lattice_class.Sup_upper) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1350 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1351 | lemma suminf_0_le: | 
| 49664 | 1352 | fixes f :: "nat \<Rightarrow> ereal" | 
| 1353 | assumes "\<And>n. 0 \<le> f n" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1354 | shows "0 \<le> (\<Sum>n. f n)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1355 | using suminf_upper[of f 0, OF assms] by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1356 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1357 | lemma suminf_le_pos: | 
| 43920 | 1358 | fixes f g :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1359 | assumes "\<And>N. f N \<le> g N" "\<And>N. 0 \<le> f N" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1360 | shows "suminf f \<le> suminf g" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1361 | proof (safe intro!: suminf_bound) | 
| 49664 | 1362 | fix n | 
| 1363 |   { fix N have "0 \<le> g N" using assms(2,1)[of N] by auto }
 | |
| 1364 |   have "setsum f {..<n} \<le> setsum g {..<n}"
 | |
| 1365 | using assms by (auto intro: setsum_mono) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1366 | also have "... \<le> suminf g" using `\<And>N. 0 \<le> g N` by (rule suminf_upper) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1367 |   finally show "setsum f {..<n} \<le> suminf g" .
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1368 | qed (rule assms(2)) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1369 | |
| 43920 | 1370 | lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal)^Suc n) = 1" | 
| 1371 | using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric] | |
| 1372 | by (simp add: one_ereal_def) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1373 | |
| 43920 | 1374 | lemma suminf_add_ereal: | 
| 1375 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1376 | assumes "\<And>i. 0 \<le> f i" "\<And>i. 0 \<le> g i" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1377 | shows "(\<Sum>i. f i + g i) = suminf f + suminf g" | 
| 43920 | 1378 | apply (subst (1 2 3) suminf_ereal_eq_SUPR) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1379 | unfolding setsum_addf | 
| 49664 | 1380 | apply (intro assms ereal_add_nonneg_nonneg SUPR_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+ | 
| 1381 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1382 | |
| 43920 | 1383 | lemma suminf_cmult_ereal: | 
| 1384 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1385 | assumes "\<And>i. 0 \<le> f i" "0 \<le> a" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1386 | shows "(\<Sum>i. a * f i) = a * suminf f" | 
| 43920 | 1387 | by (auto simp: setsum_ereal_right_distrib[symmetric] assms | 
| 1388 | ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUPR | |
| 1389 | intro!: SUPR_ereal_cmult ) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1390 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1391 | lemma suminf_PInfty: | 
| 43923 | 1392 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1393 | assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1394 | shows "f i \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1395 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1396 | from suminf_upper[of f "Suc i", OF assms(1)] assms(2) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1397 | have "(\<Sum>i<Suc i. f i) \<noteq> \<infinity>" by auto | 
| 49664 | 1398 | then show ?thesis unfolding setsum_Pinfty by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1399 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1400 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1401 | lemma suminf_PInfty_fun: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1402 | assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>" | 
| 43920 | 1403 | shows "\<exists>f'. f = (\<lambda>x. ereal (f' x))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1404 | proof - | 
| 43920 | 1405 | have "\<forall>i. \<exists>r. f i = ereal r" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1406 | proof | 
| 43920 | 1407 | fix i show "\<exists>r. f i = ereal r" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1408 | using suminf_PInfty[OF assms] assms(1)[of i] by (cases "f i") auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1409 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1410 | from choice[OF this] show ?thesis by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1411 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1412 | |
| 43920 | 1413 | lemma summable_ereal: | 
| 1414 | assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1415 | shows "summable f" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1416 | proof - | 
| 43920 | 1417 | have "0 \<le> (\<Sum>i. ereal (f i))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1418 | using assms by (intro suminf_0_le) auto | 
| 43920 | 1419 | with assms obtain r where r: "(\<Sum>i. ereal (f i)) = ereal r" | 
| 1420 | by (cases "\<Sum>i. ereal (f i)") auto | |
| 1421 | from summable_ereal_pos[of "\<lambda>x. ereal (f x)"] | |
| 1422 | have "summable (\<lambda>x. ereal (f x))" using assms by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1423 | from summable_sums[OF this] | 
| 43920 | 1424 | have "(\<lambda>x. ereal (f x)) sums (\<Sum>x. ereal (f x))" by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1425 | then show "summable f" | 
| 43920 | 1426 | unfolding r sums_ereal summable_def .. | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1427 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1428 | |
| 43920 | 1429 | lemma suminf_ereal: | 
| 1430 | assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>" | |
| 1431 | shows "(\<Sum>i. ereal (f i)) = ereal (suminf f)" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1432 | proof (rule sums_unique[symmetric]) | 
| 43920 | 1433 | from summable_ereal[OF assms] | 
| 1434 | show "(\<lambda>x. ereal (f x)) sums (ereal (suminf f))" | |
| 1435 | unfolding sums_ereal using assms by (intro summable_sums summable_ereal) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1436 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1437 | |
| 43920 | 1438 | lemma suminf_ereal_minus: | 
| 1439 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1440 | assumes ord: "\<And>i. g i \<le> f i" "\<And>i. 0 \<le> g i" and fin: "suminf f \<noteq> \<infinity>" "suminf g \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1441 | shows "(\<Sum>i. f i - g i) = suminf f - suminf g" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1442 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1443 |   { fix i have "0 \<le> f i" using ord[of i] by auto }
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1444 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1445 | from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] guess f' .. note this[simp] | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1446 | from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] guess g' .. note this[simp] | 
| 43920 | 1447 |   { fix i have "0 \<le> f i - g i" using ord[of i] by (auto simp: ereal_le_minus_iff) }
 | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1448 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1449 | have "suminf (\<lambda>i. f i - g i) \<le> suminf f" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1450 | using assms by (auto intro!: suminf_le_pos simp: field_simps) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1451 | then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>" using fin by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1452 | ultimately show ?thesis using assms `\<And>i. 0 \<le> f i` | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1453 | apply simp | 
| 49664 | 1454 | apply (subst (1 2 3) suminf_ereal) | 
| 1455 | apply (auto intro!: suminf_diff[symmetric] summable_ereal) | |
| 1456 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1457 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1458 | |
| 49664 | 1459 | lemma suminf_ereal_PInf [simp]: "(\<Sum>x. \<infinity>::ereal) = \<infinity>" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1460 | proof - | 
| 43923 | 1461 | have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>::ereal)" by (rule suminf_upper) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1462 | then show ?thesis by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1463 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1464 | |
| 43920 | 1465 | lemma summable_real_of_ereal: | 
| 43923 | 1466 | fixes f :: "nat \<Rightarrow> ereal" | 
| 49664 | 1467 | assumes f: "\<And>i. 0 \<le> f i" | 
| 1468 | and fin: "(\<Sum>i. f i) \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1469 | shows "summable (\<lambda>i. real (f i))" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1470 | proof (rule summable_def[THEN iffD2]) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1471 | have "0 \<le> (\<Sum>i. f i)" using assms by (auto intro: suminf_0_le) | 
| 43920 | 1472 | with fin obtain r where r: "ereal r = (\<Sum>i. f i)" by (cases "(\<Sum>i. f i)") auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1473 |   { fix i have "f i \<noteq> \<infinity>" using f by (intro suminf_PInfty[OF _ fin]) auto
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1474 | then have "\<bar>f i\<bar> \<noteq> \<infinity>" using f[of i] by auto } | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1475 | note fin = this | 
| 43920 | 1476 | have "(\<lambda>i. ereal (real (f i))) sums (\<Sum>i. ereal (real (f i)))" | 
| 1477 | using f by (auto intro!: summable_ereal_pos summable_sums simp: ereal_le_real_iff zero_ereal_def) | |
| 1478 | also have "\<dots> = ereal r" using fin r by (auto simp: ereal_real) | |
| 1479 | finally show "\<exists>r. (\<lambda>i. real (f i)) sums r" by (auto simp: sums_ereal) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1480 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 1481 | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1482 | lemma suminf_SUP_eq: | 
| 43920 | 1483 | fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ereal" | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1484 | assumes "\<And>i. incseq (\<lambda>n. f n i)" "\<And>n i. 0 \<le> f n i" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1485 | shows "(\<Sum>i. SUP n. f n i) = (SUP n. \<Sum>i. f n i)" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1486 | proof - | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1487 |   { fix n :: nat
 | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1488 | have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)" | 
| 43920 | 1489 | using assms by (auto intro!: SUPR_ereal_setsum[symmetric]) } | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1490 | note * = this | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1491 | show ?thesis using assms | 
| 43920 | 1492 | apply (subst (1 2) suminf_ereal_eq_SUPR) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1493 | unfolding * | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44918diff
changeset | 1494 | apply (auto intro!: SUP_upper2) | 
| 49664 | 1495 | apply (subst SUP_commute) | 
| 1496 | apply rule | |
| 1497 | done | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1498 | qed | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 1499 | |
| 47761 | 1500 | lemma suminf_setsum_ereal: | 
| 1501 | fixes f :: "_ \<Rightarrow> _ \<Rightarrow> ereal" | |
| 1502 | assumes nonneg: "\<And>i a. a \<in> A \<Longrightarrow> 0 \<le> f i a" | |
| 1503 | shows "(\<Sum>i. \<Sum>a\<in>A. f i a) = (\<Sum>a\<in>A. \<Sum>i. f i a)" | |
| 1504 | proof cases | |
| 49664 | 1505 | assume "finite A" | 
| 1506 | then show ?thesis using nonneg | |
| 47761 | 1507 | by induct (simp_all add: suminf_add_ereal setsum_nonneg) | 
| 1508 | qed simp | |
| 1509 | ||
| 50104 | 1510 | lemma suminf_ereal_eq_0: | 
| 1511 | fixes f :: "nat \<Rightarrow> ereal" | |
| 1512 | assumes nneg: "\<And>i. 0 \<le> f i" | |
| 1513 | shows "(\<Sum>i. f i) = 0 \<longleftrightarrow> (\<forall>i. f i = 0)" | |
| 1514 | proof | |
| 1515 | assume "(\<Sum>i. f i) = 0" | |
| 1516 |   { fix i assume "f i \<noteq> 0"
 | |
| 1517 | with nneg have "0 < f i" by (auto simp: less_le) | |
| 1518 | also have "f i = (\<Sum>j. if j = i then f i else 0)" | |
| 1519 |       by (subst suminf_finite[where N="{i}"]) auto
 | |
| 1520 | also have "\<dots> \<le> (\<Sum>i. f i)" | |
| 1521 | using nneg by (auto intro!: suminf_le_pos) | |
| 1522 | finally have False using `(\<Sum>i. f i) = 0` by auto } | |
| 1523 | then show "\<forall>i. f i = 0" by auto | |
| 1524 | qed simp | |
| 1525 | ||
| 44125 | 1526 | end |