| author | hoelzl | 
| Tue, 05 Mar 2013 15:43:08 +0100 | |
| changeset 51340 | 5e6296afe08d | 
| parent 51329 | 4a3c453f99a1 | 
| child 51351 | dd1dd470690b | 
| 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" | |
| 51340 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 25 | assumes "open S" shows "open (uminus ` S)" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 26 | using `open S`[unfolded open_generated_order] | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 27 | proof induct | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 28 | have "range uminus = (UNIV :: ereal set)" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 29 | by (auto simp: image_iff ereal_uminus_eq_reorder) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 30 | then show "open (range uminus :: ereal set)" by simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 31 | qed (auto simp add: image_Union image_Int) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 32 | |
| 43920 | 33 | lemma ereal_uminus_complement: | 
| 34 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 35 | shows "uminus ` (- S) = - uminus ` S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 36 | 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 | 37 | |
| 43920 | 38 | lemma ereal_closed_uminus: | 
| 39 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 40 | assumes "closed S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 41 | shows "closed (uminus ` S)" | 
| 51340 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 42 | using assms unfolding closed_def ereal_uminus_complement[symmetric] by (rule ereal_open_uminus) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 43 | |
| 43920 | 44 | lemma ereal_closed_contains_Inf: | 
| 45 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 46 |   assumes "closed S" "S ~= {}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 47 | shows "Inf S : S" | 
| 49664 | 48 | proof (rule ccontr) | 
| 49 | assume "Inf S \<notin> S" | |
| 50 | 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 | 51 | show False | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 52 | proof (cases "Inf S") | 
| 49664 | 53 | case MInf | 
| 54 | then have "(-\<infinity>) : - S" using a by auto | |
| 43920 | 55 |     then obtain y where "{..<ereal y} <= (-S)" using a open_MInfty2[of "- S"] by auto
 | 
| 49664 | 56 | 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 | 57 | 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 | 58 | then show False using MInf by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 59 | next | 
| 49664 | 60 | case PInf | 
| 61 |     then have "S={\<infinity>}" by (metis Inf_eq_PInfty assms(2))
 | |
| 44918 | 62 | 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 | 63 | next | 
| 49664 | 64 | case (real r) | 
| 65 | then have fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>" by simp | |
| 43920 | 66 | from ereal_open_cont_interval[OF a this] guess e . note e = this | 
| 49664 | 67 |     { fix x
 | 
| 68 | assume "x:S" then have "x>=Inf S" by (rule complete_lattice_class.Inf_lower) | |
| 69 | then have *: "x>Inf S-e" using e by (metis fin ereal_between(1) order_less_le_trans) | |
| 70 |       { assume "x<Inf S+e"
 | |
| 71 |         then have "x:{Inf S-e <..< Inf S+e}" using * by auto
 | |
| 72 | then have False using e `x:S` by auto | |
| 73 | } then have "x>=Inf S+e" by (metis linorder_le_less_linear) | |
| 74 | } | |
| 75 | 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 | 76 | 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 | 77 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 78 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 79 | |
| 43920 | 80 | lemma ereal_closed_contains_Sup: | 
| 81 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 82 |   assumes "closed S" "S ~= {}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 83 | shows "Sup S : S" | 
| 49664 | 84 | proof - | 
| 85 | have "closed (uminus ` S)" | |
| 86 | by (metis assms(1) ereal_closed_uminus) | |
| 87 | then have "Inf (uminus ` S) : uminus ` S" | |
| 88 | using assms ereal_closed_contains_Inf[of "uminus ` S"] by auto | |
| 89 | then have "- Sup S : uminus ` S" | |
| 90 | using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (auto simp: image_image) | |
| 91 | then show ?thesis | |
| 92 | 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 | 93 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 94 | |
| 43920 | 95 | lemma ereal_open_closed_aux: | 
| 96 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 97 | assumes "open S" "closed S" | 
| 49664 | 98 | and S: "(-\<infinity>) ~: S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 99 |   shows "S = {}"
 | 
| 49664 | 100 | proof (rule ccontr) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 101 |   assume "S ~= {}"
 | 
| 49664 | 102 | then have *: "(Inf S):S" by (metis assms(2) ereal_closed_contains_Inf) | 
| 103 |   { assume "Inf S=(-\<infinity>)"
 | |
| 104 | then have False using * assms(3) by auto } | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 105 | moreover | 
| 49664 | 106 |   { assume "Inf S=\<infinity>"
 | 
| 107 |     then have "S={\<infinity>}" by (metis Inf_eq_PInfty `S ~= {}`)
 | |
| 108 | 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 | 109 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 110 |   { assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"
 | 
| 43920 | 111 | 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 | 112 | then obtain b where b_def: "Inf S-e<b & b<Inf S" | 
| 51329 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 hoelzl parents: 
51000diff
changeset | 113 | using fin ereal_between[of "Inf S" e] dense[of "Inf S-e"] by auto | 
| 49664 | 114 |     then have "b: {Inf S-e <..< Inf S+e}" using e fin ereal_between[of "Inf S" e]
 | 
| 44918 | 115 | by auto | 
| 49664 | 116 | then have "b:S" using e by auto | 
| 117 | 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 | 118 | } ultimately show False by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 119 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 120 | |
| 43920 | 121 | lemma ereal_open_closed: | 
| 122 | fixes S :: "ereal set" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 123 |   shows "(open S & closed S) <-> (S = {} | S = UNIV)"
 | 
| 49664 | 124 | proof - | 
| 125 |   { assume lhs: "open S & closed S"
 | |
| 126 |     { assume "(-\<infinity>) ~: S"
 | |
| 127 |       then have "S={}" using lhs ereal_open_closed_aux by auto }
 | |
| 128 | moreover | |
| 129 |     { assume "(-\<infinity>) : S"
 | |
| 130 |       then have "(- S)={}" using lhs ereal_open_closed_aux[of "-S"] by auto }
 | |
| 131 |     ultimately have "S = {} | S = UNIV" by auto
 | |
| 132 | } then show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 133 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 134 | |
| 43920 | 135 | lemma ereal_open_affinity_pos: | 
| 43923 | 136 | fixes S :: "ereal set" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 137 | 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 | 138 | shows "open ((\<lambda>x. m * x + t) ` S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 139 | proof - | 
| 43920 | 140 | obtain r where r[simp]: "m = ereal r" using m by (cases m) auto | 
| 141 | 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 | 142 | have "r \<noteq> 0" "0 < r" and m': "m \<noteq> \<infinity>" "m \<noteq> -\<infinity>" "m \<noteq> 0" using m by auto | 
| 43920 | 143 | 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 | 144 | let ?f = "(\<lambda>x. m * x + t)" | 
| 49664 | 145 | show ?thesis | 
| 146 | unfolding open_ereal_def | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 147 | proof (intro conjI impI exI subsetI) | 
| 43920 | 148 | 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 | 149 | proof safe | 
| 49664 | 150 | fix x y | 
| 151 | assume "ereal y = m * x + t" "x \<in> S" | |
| 43920 | 152 | 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 | 153 | 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 | 154 | qed force | 
| 43920 | 155 | then show "open (ereal -` ?f ` S)" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 156 | 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 | 157 | next | 
| 49664 | 158 | assume "\<infinity> \<in> ?f`S" | 
| 159 | with `0 < r` have "\<infinity> \<in> S" by auto | |
| 160 | fix x | |
| 161 |     assume "x \<in> {ereal (r * l + p)<..}"
 | |
| 43920 | 162 | 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 | 163 | show "x \<in> ?f`S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 164 | proof (rule image_eqI) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 165 | show "x = m * ((x - t) / m) + t" | 
| 43920 | 166 | using m t by (cases rule: ereal3_cases[of m x t]) auto | 
| 167 | have "ereal l < (x - t)/m" | |
| 168 | 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 | 169 | 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 | 170 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 171 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 172 | assume "-\<infinity> \<in> ?f`S" with `0 < r` have "-\<infinity> \<in> S" by auto | 
| 43920 | 173 |     fix x assume "x \<in> {..<ereal (r * u + p)}"
 | 
| 174 | 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 | 175 | show "x \<in> ?f`S" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 176 | proof (rule image_eqI) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 177 | show "x = m * ((x - t) / m) + t" | 
| 43920 | 178 | using m t by (cases rule: ereal3_cases[of m x t]) auto | 
| 179 | have "(x - t)/m < ereal u" | |
| 180 | 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 | 181 | 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 | 182 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 183 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 184 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 185 | |
| 43920 | 186 | lemma ereal_open_affinity: | 
| 43923 | 187 | fixes S :: "ereal set" | 
| 49664 | 188 | assumes "open S" | 
| 189 | and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0" | |
| 190 | and t: "\<bar>t\<bar> \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 191 | shows "open ((\<lambda>x. m * x + t) ` S)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 192 | proof cases | 
| 49664 | 193 | assume "0 < m" | 
| 194 | then show ?thesis | |
| 43920 | 195 | 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 | 196 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 197 | assume "\<not> 0 < m" then | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 198 | 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 | 199 | then have m: "-m \<noteq> \<infinity>" "0 < -m" using `\<bar>m\<bar> \<noteq> \<infinity>` | 
| 43920 | 200 | by (auto simp: ereal_uminus_eq_reorder) | 
| 201 | 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 | 202 | show ?thesis unfolding image_image by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 203 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 204 | |
| 43920 | 205 | lemma ereal_lim_mult: | 
| 206 | fixes X :: "'a \<Rightarrow> ereal" | |
| 49664 | 207 | assumes lim: "(X ---> L) net" | 
| 208 | and a: "\<bar>a\<bar> \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 209 | shows "((\<lambda>i. a * X i) ---> a * L) net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 210 | proof cases | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 211 | assume "a \<noteq> 0" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 212 | show ?thesis | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 213 | proof (rule topological_tendstoI) | 
| 49664 | 214 | fix S | 
| 215 | assume "open S" "a * L \<in> S" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 216 | have "a * L / a = L" | 
| 43920 | 217 | 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 | 218 | 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 | 219 | 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 | 220 | moreover have "open ((\<lambda>x. x / a) ` S)" | 
| 43920 | 221 | using ereal_open_affinity[OF `open S`, of "inverse a" 0] `a \<noteq> 0` a | 
| 222 | 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 | 223 | note * = lim[THEN topological_tendstoD, OF this L] | 
| 49664 | 224 |     { fix x
 | 
| 225 | from a `a \<noteq> 0` have "a * (x / a) = x" | |
| 43920 | 226 | 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 | 227 | note this[simp] | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 228 | 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 | 229 | by (rule eventually_mono[OF _ *]) auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 230 | qed | 
| 44918 | 231 | qed auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 232 | |
| 43920 | 233 | lemma ereal_lim_uminus: | 
| 49664 | 234 | fixes X :: "'a \<Rightarrow> ereal" | 
| 235 | shows "((\<lambda>i. - X i) ---> -L) net \<longleftrightarrow> (X ---> L) net" | |
| 43920 | 236 | using ereal_lim_mult[of X L net "ereal (-1)"] | 
| 49664 | 237 | 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 | 238 | by (auto simp add: algebra_simps) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 239 | |
| 43920 | 240 | lemma Lim_bounded2_ereal: | 
| 241 | assumes lim:"f ----> (l :: ereal)" | |
| 49664 | 242 | and ge: "ALL n>=N. f n >= C" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 243 | shows "l>=C" | 
| 51000 | 244 | using ge | 
| 245 | by (intro tendsto_le[OF trivial_limit_sequentially lim tendsto_const]) | |
| 246 | (auto simp: eventually_sequentially) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 247 | |
| 43923 | 248 | 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 | 249 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 250 |   assume "x = -\<infinity>" then have "{x..} = UNIV" by auto
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 251 |   then show "open {x..}" by auto
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 252 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 253 |   assume "open {x..}"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 254 |   then have "open {x..} \<and> closed {x..}" by auto
 | 
| 43920 | 255 |   then have "{x..} = UNIV" unfolding ereal_open_closed by auto
 | 
| 256 | 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 | 257 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 258 | |
| 43920 | 259 | lemma open_uminus_iff: "open (uminus ` S) \<longleftrightarrow> open (S::ereal set)" | 
| 260 | 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 | 261 | |
| 43920 | 262 | lemma ereal_Liminf_uminus: | 
| 263 | fixes f :: "'a => ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 264 | shows "Liminf net (\<lambda>x. - (f x)) = -(Limsup net f)" | 
| 43920 | 265 | 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 | 266 | |
| 43920 | 267 | lemma ereal_Lim_uminus: | 
| 49664 | 268 | fixes f :: "'a \<Rightarrow> ereal" | 
| 269 | 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 | 270 | using | 
| 43920 | 271 | ereal_lim_mult[of f f0 net "- 1"] | 
| 272 | ereal_lim_mult[of "\<lambda>x. - (f x)" "-f0" net "- 1"] | |
| 273 | by (auto simp: ereal_uminus_reorder) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 274 | |
| 50104 | 275 | lemma convergent_ereal_limsup: | 
| 276 | fixes X :: "nat \<Rightarrow> ereal" | |
| 277 | shows "convergent X \<Longrightarrow> limsup X = lim X" | |
| 278 | by (auto simp: convergent_def limI lim_imp_Limsup) | |
| 279 | ||
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 280 | lemma Liminf_PInfty: | 
| 43920 | 281 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 282 | assumes "\<not> trivial_limit net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 283 | shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 284 | proof (intro lim_imp_Liminf iffI assms) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 285 | assume rhs: "Liminf net f = \<infinity>" | 
| 51000 | 286 | show "(f ---> \<infinity>) net" | 
| 287 | unfolding tendsto_PInfty | |
| 288 | proof | |
| 289 | fix r :: real | |
| 290 | have "ereal r < top" unfolding top_ereal_def by simp | |
| 291 | with rhs obtain P where "eventually P net" "r < INFI (Collect P) f" | |
| 292 | unfolding Liminf_def SUP_eq_top_iff top_ereal_def[symmetric] by auto | |
| 293 | then show "eventually (\<lambda>x. ereal r < f x) net" | |
| 294 | by (auto elim!: eventually_elim1 dest: less_INF_D) | |
| 295 | qed | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 296 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 297 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 298 | lemma Limsup_MInfty: | 
| 43920 | 299 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 300 | assumes "\<not> trivial_limit net" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 301 | shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>" | 
| 43920 | 302 | using assms ereal_Lim_uminus[of f "-\<infinity>"] Liminf_PInfty[of _ "\<lambda>x. - (f x)"] | 
| 303 | 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 | 304 | |
| 50104 | 305 | lemma convergent_ereal: | 
| 306 | fixes X :: "nat \<Rightarrow> ereal" | |
| 307 | shows "convergent X \<longleftrightarrow> limsup X = liminf X" | |
| 51340 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 308 | using tendsto_iff_Liminf_eq_Limsup[of sequentially] | 
| 50104 | 309 | by (auto simp: convergent_def) | 
| 310 | ||
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 311 | lemma limsup_INFI_SUPR: | 
| 43920 | 312 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 313 |   shows "limsup f = (INF n. SUP m:{n..}. f m)"
 | 
| 43920 | 314 | using ereal_Limsup_uminus[of sequentially "\<lambda>x. - f x"] | 
| 315 | 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 | 316 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 317 | lemma liminf_PInfty: | 
| 43920 | 318 | fixes X :: "nat => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 319 | shows "X ----> \<infinity> <-> liminf X = \<infinity>" | 
| 49664 | 320 | by (metis Liminf_PInfty trivial_limit_sequentially) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 321 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 322 | lemma limsup_MInfty: | 
| 43920 | 323 | fixes X :: "nat => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 324 | shows "X ----> (-\<infinity>) <-> limsup X = (-\<infinity>)" | 
| 49664 | 325 | by (metis Limsup_MInfty trivial_limit_sequentially) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 326 | |
| 43920 | 327 | lemma ereal_lim_mono: | 
| 328 | fixes X Y :: "nat => ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 329 | assumes "\<And>n. N \<le> n \<Longrightarrow> X n <= Y n" | 
| 49664 | 330 | and "X ----> x" "Y ----> y" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 331 | shows "x <= y" | 
| 51000 | 332 | using assms(1) by (intro LIMSEQ_le[OF assms(2,3)]) auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 333 | |
| 43920 | 334 | lemma incseq_le_ereal: | 
| 335 | fixes X :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 336 | assumes inc: "incseq X" and lim: "X ----> L" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 337 | shows "X N \<le> L" | 
| 49664 | 338 | 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 | 339 | |
| 49664 | 340 | lemma decseq_ge_ereal: | 
| 341 | assumes dec: "decseq X" | |
| 342 | and lim: "X ----> (L::ereal)" | |
| 343 | shows "X N >= L" | |
| 344 | 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 | 345 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 346 | lemma bounded_abs: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 347 | assumes "(a::real)<=x" "x<=b" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 348 | shows "abs x <= max (abs a) (abs b)" | 
| 49664 | 349 | by (metis abs_less_iff assms leI le_max_iff_disj | 
| 350 | 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 | 351 | |
| 49664 | 352 | lemma lim_ereal_increasing: | 
| 51000 | 353 | assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<ge> f m" | 
| 354 |   obtains l where "f ----> (l::'a::{complete_linorder, linorder_topology})"
 | |
| 355 | proof | |
| 356 | show "f ----> (SUP n. f n)" | |
| 357 | using assms | |
| 358 | by (intro increasing_tendsto) | |
| 359 | (auto simp: SUP_upper eventually_sequentially less_SUP_iff intro: less_le_trans) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 360 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 361 | |
| 49664 | 362 | lemma lim_ereal_decreasing: | 
| 51000 | 363 | assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<le> f m" | 
| 364 |   obtains l where "f ----> (l::'a::{complete_linorder, linorder_topology})"
 | |
| 365 | proof | |
| 366 | show "f ----> (INF n. f n)" | |
| 367 | using assms | |
| 368 | by (intro decreasing_tendsto) | |
| 369 | (auto simp: INF_lower eventually_sequentially INF_less_iff intro: le_less_trans) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 370 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 371 | |
| 43920 | 372 | lemma compact_ereal: | 
| 373 | fixes X :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 374 | 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 | 375 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 376 | 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 | 377 | 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 | 378 | 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 | 379 | by (auto simp add: monoseq_def) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 380 | then obtain l where "(X\<circ>r) ----> l" | 
| 43920 | 381 | 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 | 382 | then show ?thesis using `subseq r` by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 383 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 384 | |
| 43920 | 385 | lemma ereal_Sup_lim: | 
| 386 | 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 | 387 | shows "a \<le> Sup s" | 
| 49664 | 388 | 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 | 389 | |
| 43920 | 390 | lemma ereal_Inf_lim: | 
| 391 | 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 | 392 | shows "Inf s \<le> a" | 
| 49664 | 393 | 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 | 394 | |
| 43920 | 395 | lemma SUP_Lim_ereal: | 
| 51000 | 396 |   fixes X :: "nat \<Rightarrow> 'a::{complete_linorder, linorder_topology}"
 | 
| 397 | assumes inc: "incseq X" and l: "X ----> l" | |
| 49664 | 398 | shows "(SUP n. X n) = l" | 
| 51000 | 399 | using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l] by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 400 | |
| 43920 | 401 | lemma INF_Lim_ereal: "decseq X \<Longrightarrow> X ----> l \<Longrightarrow> (INF n. X n) = (l::ereal)" | 
| 402 | using SUP_Lim_ereal[of "\<lambda>i. - X i" "- l"] | |
| 403 | 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 | 404 | |
| 43920 | 405 | lemma LIMSEQ_ereal_INFI: "decseq X \<Longrightarrow> X ----> (INF n. X n :: ereal)" | 
| 51000 | 406 | using LIMSEQ_SUP[of "\<lambda>i. - X i"] | 
| 43920 | 407 | 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 | 408 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 409 | lemma SUP_eq_LIMSEQ: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 410 | assumes "mono f" | 
| 43920 | 411 | 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 | 412 | proof | 
| 43920 | 413 | have inc: "incseq (\<lambda>i. ereal (f i))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 414 | 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 | 415 |   { assume "f ----> x"
 | 
| 49664 | 416 | then have "(\<lambda>i. ereal (f i)) ----> ereal x" by auto | 
| 417 | from SUP_Lim_ereal[OF inc this] | |
| 418 | show "(SUP n. ereal (f n)) = ereal x" . } | |
| 43920 | 419 |   { assume "(SUP n. ereal (f n)) = ereal x"
 | 
| 51000 | 420 | with LIMSEQ_SUP[OF inc] | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 421 | show "f ----> x" by auto } | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 422 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 423 | |
| 49664 | 424 | lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
 | 
| 425 | unfolding islimpt_def by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 426 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 427 | |
| 49664 | 428 | lemma islimpt_in_closure: "(x islimpt S) = (x:closure(S-{x}))"
 | 
| 429 | unfolding closure_def using islimpt_punctured by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 430 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 431 | |
| 49664 | 432 | lemma not_trivial_limit_within: "~trivial_limit (at x within S) = (x:closure(S-{x}))"
 | 
| 433 | 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 | 434 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 435 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 436 | lemma not_trivial_limit_within_ball: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 437 |   "(~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 | 438 | (is "?lhs = ?rhs") | 
| 49664 | 439 | proof - | 
| 440 |   { assume "?lhs"
 | |
| 441 |     { fix e :: real
 | |
| 442 | assume "e>0" | |
| 443 |       then obtain y where "y:(S-{x}) & dist y x < e"
 | |
| 444 |         using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
 | |
| 445 | by auto | |
| 446 |       then have "y : (S Int ball x e - {x})"
 | |
| 447 | unfolding ball_def by (simp add: dist_commute) | |
| 448 |       then have "S Int ball x e - {x} ~= {}" by blast
 | |
| 449 | } then have "?rhs" by auto | |
| 450 | } | |
| 451 | moreover | |
| 452 |   { assume "?rhs"
 | |
| 453 |     { fix e :: real
 | |
| 454 | assume "e>0" | |
| 455 |       then obtain y where "y : (S Int ball x e - {x})" using `?rhs` by blast
 | |
| 456 |       then have "y:(S-{x}) & dist y x < e"
 | |
| 457 | unfolding ball_def by (simp add: dist_commute) | |
| 458 |       then have "EX y:(S-{x}). dist y x < e" by auto
 | |
| 459 | } | |
| 460 | then have "?lhs" | |
| 461 |       using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"] by auto
 | |
| 462 | } | |
| 463 | ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 464 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 465 | |
| 43920 | 466 | lemma liminf_ereal_cminus: | 
| 49664 | 467 | fixes f :: "nat \<Rightarrow> ereal" | 
| 468 | assumes "c \<noteq> -\<infinity>" | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 469 | shows "liminf (\<lambda>x. c - f x) = c - limsup f" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 470 | proof (cases c) | 
| 49664 | 471 | case PInf | 
| 472 | then show ?thesis by (simp add: Liminf_const) | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 473 | next | 
| 49664 | 474 | case (real r) | 
| 475 | then show ?thesis | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 476 | unfolding liminf_SUPR_INFI limsup_INFI_SUPR | 
| 43920 | 477 | apply (subst INFI_ereal_cminus) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 478 | apply auto | 
| 43920 | 479 | apply (subst SUPR_ereal_cminus) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 480 | apply auto | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 481 | done | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 482 | qed (insert `c \<noteq> -\<infinity>`, simp) | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 483 | |
| 49664 | 484 | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 485 | subsubsection {* Continuity *}
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 486 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 487 | lemma continuous_imp_tendsto: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 488 | assumes "continuous (at x0) f" | 
| 49664 | 489 | and "x ----> x0" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 490 | shows "(f o x) ----> (f x0)" | 
| 49664 | 491 | proof - | 
| 492 |   { fix S
 | |
| 493 | assume "open S & (f x0):S" | |
| 494 | then obtain T where T_def: "open T & x0 : T & (ALL x:T. f x : S)" | |
| 495 | using assms continuous_at_open by metis | |
| 496 | then have "(EX N. ALL n>=N. x n : T)" | |
| 497 | using assms tendsto_explicit T_def by auto | |
| 498 | then have "(EX N. ALL n>=N. f(x n) : S)" using T_def by auto | |
| 499 | } | |
| 500 | 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 | 501 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 502 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 503 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 504 | lemma continuous_at_sequentially2: | 
| 49664 | 505 | fixes f :: "'a::metric_space => 'b:: topological_space" | 
| 506 | shows "continuous (at x0) f <-> (ALL x. (x ----> x0) --> (f o x) ----> (f x0))" | |
| 507 | proof - | |
| 508 |   { assume "~(continuous (at x0) f)"
 | |
| 509 | then obtain T where | |
| 510 | T_def: "open T & f x0 : T & (ALL S. (open S & x0 : S) --> (EX x':S. f x' ~: T))" | |
| 511 | using continuous_at_open[of x0 f] by metis | |
| 512 |     def X == "{x'. f x' ~: T}"
 | |
| 513 | then have "x0 islimpt X" | |
| 514 | unfolding islimpt_def using T_def by auto | |
| 515 | then obtain x where x_def: "(ALL n. x n : X) & x ----> x0" | |
| 516 | using islimpt_sequential[of x0 X] by auto | |
| 517 | then have "~(f o x) ----> (f x0)" | |
| 518 | unfolding tendsto_explicit using X_def T_def by auto | |
| 519 | then have "EX x. x ----> x0 & (~(f o x) ----> (f x0))" using x_def by auto | |
| 520 | } | |
| 521 | 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 | 522 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 523 | |
| 43920 | 524 | lemma continuous_at_of_ereal: | 
| 525 | fixes x0 :: ereal | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 526 | assumes "\<bar>x0\<bar> \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 527 | shows "continuous (at x0) real" | 
| 49664 | 528 | proof - | 
| 529 |   { fix T
 | |
| 530 | assume T_def: "open T & real x0 : T" | |
| 531 | def S == "ereal ` T" | |
| 532 | then have "ereal (real x0) : S" using T_def by auto | |
| 533 | then have "x0 : S" using assms ereal_real by auto | |
| 534 | moreover have "open S" using open_ereal S_def T_def by auto | |
| 535 | moreover have "ALL y:S. real y : T" using S_def T_def by auto | |
| 536 | ultimately have "EX S. x0 : S & open S & (ALL y:S. real y : T)" by auto | |
| 537 | } | |
| 538 | 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 | 539 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 540 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 541 | |
| 43920 | 542 | lemma continuous_at_iff_ereal: | 
| 49664 | 543 | fixes f :: "'a::t2_space => real" | 
| 544 | shows "continuous (at x0) f <-> continuous (at x0) (ereal o f)" | |
| 545 | proof - | |
| 546 |   { assume "continuous (at x0) f"
 | |
| 547 | then have "continuous (at x0) (ereal o f)" | |
| 548 | using continuous_at_ereal continuous_at_compose[of x0 f ereal] by auto | |
| 549 | } | |
| 550 | moreover | |
| 551 |   { assume "continuous (at x0) (ereal o f)"
 | |
| 552 | then have "continuous (at x0) (real o (ereal o f))" | |
| 553 | using continuous_at_of_ereal by (intro continuous_at_compose[of x0 "ereal o f"]) auto | |
| 554 | moreover have "real o (ereal o f) = f" using real_ereal_id by (simp add: o_assoc) | |
| 555 | ultimately have "continuous (at x0) f" by auto | |
| 556 | } ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 557 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 558 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 559 | |
| 43920 | 560 | lemma continuous_on_iff_ereal: | 
| 49664 | 561 | fixes f :: "'a::t2_space => real" | 
| 562 | fixes A assumes "open A" | |
| 563 | shows "continuous_on A f <-> continuous_on A (ereal o f)" | |
| 564 | 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 | 565 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 566 | |
| 43923 | 567 | lemma continuous_on_real: "continuous_on (UNIV-{\<infinity>,(-\<infinity>::ereal)}) real"
 | 
| 49664 | 568 | 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 | 569 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 570 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 571 | lemma continuous_on_iff_real: | 
| 43920 | 572 | fixes f :: "'a::t2_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 573 | 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 | 574 | shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)" | 
| 49664 | 575 | proof - | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 576 |   have "f ` A <= UNIV-{\<infinity>,(-\<infinity>)}" using assms by force
 | 
| 49664 | 577 | then have *: "continuous_on (f ` A) real" | 
| 578 | using continuous_on_real by (simp add: continuous_on_subset) | |
| 579 | have **: "continuous_on ((real o f) ` A) ereal" | |
| 580 | using continuous_on_ereal continuous_on_subset[of "UNIV" "ereal" "(real o f) ` A"] by blast | |
| 581 |   { assume "continuous_on A f"
 | |
| 582 | then have "continuous_on A (real o f)" | |
| 583 | apply (subst continuous_on_compose) | |
| 584 | using * apply auto | |
| 585 | done | |
| 586 | } | |
| 587 | moreover | |
| 588 |   { assume "continuous_on A (real o f)"
 | |
| 589 | then have "continuous_on A (ereal o (real o f))" | |
| 590 | apply (subst continuous_on_compose) | |
| 591 | using ** apply auto | |
| 592 | done | |
| 593 | then have "continuous_on A f" | |
| 594 | apply (subst continuous_on_eq[of A "ereal o (real o f)" f]) | |
| 595 | using assms ereal_real apply auto | |
| 596 | done | |
| 597 | } | |
| 598 | ultimately show ?thesis by auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 599 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 600 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 601 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 602 | lemma continuous_at_const: | 
| 43920 | 603 | fixes f :: "'a::t2_space => ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 604 | assumes "ALL x. (f x = C)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 605 | shows "ALL x. continuous (at x) f" | 
| 49664 | 606 | 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 | 607 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 608 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 609 | lemma closure_contains_Inf: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 610 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 611 |   assumes "S ~= {}" "EX B. ALL x:S. B<=x"
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 612 | shows "Inf S : closure S" | 
| 49664 | 613 | proof - | 
| 614 | have *: "ALL x:S. Inf S <= x" | |
| 615 | using Inf_lower_EX[of _ S] assms by metis | |
| 616 |   { fix e
 | |
| 617 | assume "e>(0 :: real)" | |
| 618 |     then obtain x where x_def: "x:S & x < Inf S + e" using Inf_close `S ~= {}` by auto
 | |
| 619 | moreover then have "x > Inf S - e" using * by auto | |
| 620 | ultimately have "abs (x - Inf S) < e" by (simp add: abs_diff_less_iff) | |
| 621 | then have "EX x:S. abs (x - Inf S) < e" using x_def by auto | |
| 622 | } | |
| 623 | then show ?thesis | |
| 624 | apply (subst closure_approachable) | |
| 625 | unfolding dist_norm apply auto | |
| 626 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 627 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 628 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 629 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 630 | lemma closed_contains_Inf: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 631 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 632 |   assumes "S ~= {}" "EX B. ALL x:S. B<=x"
 | 
| 49664 | 633 | and "closed S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 634 | shows "Inf S : S" | 
| 49664 | 635 | by (metis closure_contains_Inf closure_closed assms) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 636 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 637 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 638 | lemma mono_closed_real: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 639 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 640 | assumes mono: "ALL y z. y:S & y<=z --> z:S" | 
| 49664 | 641 | and "closed S" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 642 |   shows "S = {} | S = UNIV | (EX a. S = {a ..})"
 | 
| 49664 | 643 | proof - | 
| 644 |   { assume "S ~= {}"
 | |
| 645 |     { assume ex: "EX B. ALL x:S. B<=x"
 | |
| 646 | then have *: "ALL x:S. Inf S <= x" using Inf_lower_EX[of _ S] ex by metis | |
| 647 |       then have "Inf S : S" apply (subst closed_contains_Inf) using ex `S ~= {}` `closed S` by auto
 | |
| 648 | then have "ALL x. (Inf S <= x <-> x:S)" using mono[rule_format, of "Inf S"] * by auto | |
| 649 |       then have "S = {Inf S ..}" by auto
 | |
| 650 |       then have "EX a. S = {a ..}" by auto
 | |
| 651 | } | |
| 652 | moreover | |
| 653 |     { assume "~(EX B. ALL x:S. B<=x)"
 | |
| 654 | then have nex: "ALL B. EX x:S. x<B" by (simp add: not_le) | |
| 655 |       { fix y
 | |
| 656 | obtain x where "x:S & x < y" using nex by auto | |
| 657 | then have "y:S" using mono[rule_format, of x y] by auto | |
| 658 | } then have "S = UNIV" by auto | |
| 659 | } | |
| 660 |     ultimately have "S = UNIV | (EX a. S = {a ..})" by blast
 | |
| 661 | } then show ?thesis by blast | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 662 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 663 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 664 | |
| 43920 | 665 | lemma mono_closed_ereal: | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 666 | fixes S :: "real set" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 667 | assumes mono: "ALL y z. y:S & y<=z --> z:S" | 
| 49664 | 668 | and "closed S" | 
| 43920 | 669 |   shows "EX a. S = {x. a <= ereal x}"
 | 
| 49664 | 670 | proof - | 
| 671 |   { assume "S = {}"
 | |
| 672 | then have ?thesis apply(rule_tac x=PInfty in exI) by auto } | |
| 673 | moreover | |
| 674 |   { assume "S = UNIV"
 | |
| 675 | then have ?thesis apply(rule_tac x="-\<infinity>" in exI) by auto } | |
| 676 | moreover | |
| 677 |   { assume "EX a. S = {a ..}"
 | |
| 678 |     then obtain a where "S={a ..}" by auto
 | |
| 679 | then have ?thesis apply(rule_tac x="ereal a" in exI) by auto | |
| 680 | } | |
| 681 | 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 | 682 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 683 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 684 | subsection {* Sums *}
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 685 | |
| 49664 | 686 | 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 | 687 | proof cases | 
| 49664 | 688 | assume "finite A" | 
| 689 | then show ?thesis by induct auto | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 690 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 691 | |
| 43923 | 692 | lemma setsum_Pinfty: | 
| 693 | fixes f :: "'a \<Rightarrow> ereal" | |
| 694 | 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 | 695 | proof safe | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 696 | assume *: "setsum f P = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 697 | show "finite P" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 698 | 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 | 699 | show "\<exists>i\<in>P. f i = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 700 | proof (rule ccontr) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 701 | 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 | 702 | 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 | 703 | by induct auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 704 | with * show False by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 705 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 706 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 707 | fix i assume "finite P" "i \<in> P" "f i = \<infinity>" | 
| 49664 | 708 | then show "setsum f P = \<infinity>" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 709 | proof induct | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 710 | case (insert x A) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 711 | show ?case using insert by (cases "x = i") auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 712 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 713 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 714 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 715 | lemma setsum_Inf: | 
| 43923 | 716 | fixes f :: "'a \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 717 | 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 | 718 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 719 | assume *: "\<bar>setsum f A\<bar> = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 720 | have "finite A" by (rule ccontr) (insert *, auto) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 721 | 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 | 722 | proof (rule ccontr) | 
| 43920 | 723 | 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 | 724 | from bchoice[OF this] guess r .. | 
| 44142 | 725 | with * show False by auto | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 726 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 727 | 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 | 728 | next | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 729 | 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 | 730 | 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 | 731 | then show "\<bar>setsum f A\<bar> = \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 732 | proof induct | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 733 | case (insert j A) then show ?case | 
| 43920 | 734 | 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 | 735 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 736 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 737 | |
| 43920 | 738 | lemma setsum_real_of_ereal: | 
| 43923 | 739 | fixes f :: "'i \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 740 | 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 | 741 | 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 | 742 | proof - | 
| 43920 | 743 | 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 | 744 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 745 | fix x assume "x \<in> S" | 
| 43920 | 746 | 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 | 747 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 748 | from bchoice[OF this] guess r .. | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 749 | then show ?thesis by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 750 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 751 | |
| 43920 | 752 | lemma setsum_ereal_0: | 
| 753 | 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 | 754 | 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 | 755 | proof | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 756 | assume *: "(\<Sum>x\<in>A. f x) = 0" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 757 | 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 | 758 | then have "\<forall>i\<in>A. \<bar>f i\<bar> \<noteq> \<infinity>" using assms by (force simp: setsum_Pinfty) | 
| 43920 | 759 | 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 | 760 | 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 | 761 | 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 | 762 | qed (rule setsum_0') | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 763 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 764 | |
| 43920 | 765 | lemma setsum_ereal_right_distrib: | 
| 49664 | 766 | fixes f :: "'a \<Rightarrow> ereal" | 
| 767 | 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 | 768 | 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 | 769 | proof cases | 
| 49664 | 770 | assume "finite A" | 
| 771 | then show ?thesis using assms | |
| 43920 | 772 | 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 | 773 | qed simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 774 | |
| 43920 | 775 | lemma sums_ereal_positive: | 
| 49664 | 776 | fixes f :: "nat \<Rightarrow> ereal" | 
| 777 | assumes "\<And>i. 0 \<le> f i" | |
| 778 | 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 | 779 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 780 | have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)" | 
| 43920 | 781 | using ereal_add_mono[OF _ assms] by (auto intro!: incseq_SucI) | 
| 51000 | 782 | from LIMSEQ_SUP[OF this] | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 783 | show ?thesis unfolding sums_def by (simp add: atLeast0LessThan) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 784 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 785 | |
| 43920 | 786 | lemma summable_ereal_pos: | 
| 49664 | 787 | fixes f :: "nat \<Rightarrow> ereal" | 
| 788 | assumes "\<And>i. 0 \<le> f i" | |
| 789 | shows "summable f" | |
| 43920 | 790 | 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 | 791 | |
| 43920 | 792 | lemma suminf_ereal_eq_SUPR: | 
| 49664 | 793 | fixes f :: "nat \<Rightarrow> ereal" | 
| 794 | assumes "\<And>i. 0 \<le> f i" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 795 | shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)" | 
| 43920 | 796 | 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 | 797 | |
| 49664 | 798 | 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 | 799 | unfolding sums_def by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 800 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 801 | lemma suminf_bound: | 
| 43920 | 802 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 803 | 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 | 804 | shows "suminf f \<le> x" | 
| 43920 | 805 | proof (rule Lim_bounded_ereal) | 
| 806 | 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 | 807 | 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 | 808 | by (auto dest!: summable_sums simp: sums_def atLeast0LessThan) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 809 |   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 | 810 | using assms by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 811 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 812 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 813 | lemma suminf_bound_add: | 
| 43920 | 814 | fixes f :: "nat \<Rightarrow> ereal" | 
| 49664 | 815 | assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x" | 
| 816 | and pos: "\<And>n. 0 \<le> f n" | |
| 817 | and "y \<noteq> -\<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 818 | shows "suminf f + y \<le> x" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 819 | proof (cases y) | 
| 49664 | 820 | case (real r) | 
| 821 | then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" | |
| 43920 | 822 | using assms by (simp add: ereal_le_minus) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 823 | 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 | 824 | then show "(\<Sum> n. f n) + y \<le> x" | 
| 43920 | 825 | 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 | 826 | qed (insert assms, auto) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 827 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 828 | lemma suminf_upper: | 
| 49664 | 829 | fixes f :: "nat \<Rightarrow> ereal" | 
| 830 | assumes "\<And>n. 0 \<le> f n" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 831 | shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)" | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44918diff
changeset | 832 | unfolding suminf_ereal_eq_SUPR[OF assms] SUP_def | 
| 45031 | 833 | by (auto intro: complete_lattice_class.Sup_upper) | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 834 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 835 | lemma suminf_0_le: | 
| 49664 | 836 | fixes f :: "nat \<Rightarrow> ereal" | 
| 837 | assumes "\<And>n. 0 \<le> f n" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 838 | shows "0 \<le> (\<Sum>n. f n)" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 839 | 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 | 840 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 841 | lemma suminf_le_pos: | 
| 43920 | 842 | fixes f g :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 843 | 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 | 844 | shows "suminf f \<le> suminf g" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 845 | proof (safe intro!: suminf_bound) | 
| 49664 | 846 | fix n | 
| 847 |   { fix N have "0 \<le> g N" using assms(2,1)[of N] by auto }
 | |
| 848 |   have "setsum f {..<n} \<le> setsum g {..<n}"
 | |
| 849 | using assms by (auto intro: setsum_mono) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 850 | 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 | 851 |   finally show "setsum f {..<n} \<le> suminf g" .
 | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 852 | qed (rule assms(2)) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 853 | |
| 43920 | 854 | lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal)^Suc n) = 1" | 
| 855 | using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric] | |
| 856 | by (simp add: one_ereal_def) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 857 | |
| 43920 | 858 | lemma suminf_add_ereal: | 
| 859 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 860 | 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 | 861 | shows "(\<Sum>i. f i + g i) = suminf f + suminf g" | 
| 43920 | 862 | 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 | 863 | unfolding setsum_addf | 
| 49664 | 864 | apply (intro assms ereal_add_nonneg_nonneg SUPR_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+ | 
| 865 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 866 | |
| 43920 | 867 | lemma suminf_cmult_ereal: | 
| 868 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 869 | 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 | 870 | shows "(\<Sum>i. a * f i) = a * suminf f" | 
| 43920 | 871 | by (auto simp: setsum_ereal_right_distrib[symmetric] assms | 
| 872 | ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUPR | |
| 873 | intro!: SUPR_ereal_cmult ) | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 874 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 875 | lemma suminf_PInfty: | 
| 43923 | 876 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 877 | 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 | 878 | shows "f i \<noteq> \<infinity>" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 879 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 880 | 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 | 881 | have "(\<Sum>i<Suc i. f i) \<noteq> \<infinity>" by auto | 
| 49664 | 882 | then show ?thesis unfolding setsum_Pinfty by simp | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 883 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 884 | |
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 885 | lemma suminf_PInfty_fun: | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 886 | assumes "\<And>i. 0 \<le> f i" "suminf f \<noteq> \<infinity>" | 
| 43920 | 887 | 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 | 888 | proof - | 
| 43920 | 889 | 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 | 890 | proof | 
| 43920 | 891 | 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 | 892 | 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 | 893 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 894 | from choice[OF this] show ?thesis by auto | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 895 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 896 | |
| 43920 | 897 | lemma summable_ereal: | 
| 898 | 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 | 899 | shows "summable f" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 900 | proof - | 
| 43920 | 901 | have "0 \<le> (\<Sum>i. ereal (f i))" | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 902 | using assms by (intro suminf_0_le) auto | 
| 43920 | 903 | with assms obtain r where r: "(\<Sum>i. ereal (f i)) = ereal r" | 
| 904 | by (cases "\<Sum>i. ereal (f i)") auto | |
| 905 | from summable_ereal_pos[of "\<lambda>x. ereal (f x)"] | |
| 906 | 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 | 907 | from summable_sums[OF this] | 
| 43920 | 908 | 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 | 909 | then show "summable f" | 
| 43920 | 910 | unfolding r sums_ereal summable_def .. | 
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 911 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 912 | |
| 43920 | 913 | lemma suminf_ereal: | 
| 914 | assumes "\<And>i. 0 \<le> f i" "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>" | |
| 915 | 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 | 916 | proof (rule sums_unique[symmetric]) | 
| 43920 | 917 | from summable_ereal[OF assms] | 
| 918 | show "(\<lambda>x. ereal (f x)) sums (ereal (suminf f))" | |
| 919 | 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 | 920 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 921 | |
| 43920 | 922 | lemma suminf_ereal_minus: | 
| 923 | fixes f g :: "nat \<Rightarrow> ereal" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 924 | 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 | 925 | 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 | 926 | proof - | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 927 |   { 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 | 928 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 929 | 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 | 930 | from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] guess g' .. note this[simp] | 
| 43920 | 931 |   { 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 | 932 | moreover | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 933 | 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 | 934 | 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 | 935 | 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 | 936 | 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 | 937 | apply simp | 
| 49664 | 938 | apply (subst (1 2 3) suminf_ereal) | 
| 939 | apply (auto intro!: suminf_diff[symmetric] summable_ereal) | |
| 940 | done | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 941 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 942 | |
| 49664 | 943 | 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 | 944 | proof - | 
| 43923 | 945 | 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 | 946 | then show ?thesis by simp | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 947 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 948 | |
| 43920 | 949 | lemma summable_real_of_ereal: | 
| 43923 | 950 | fixes f :: "nat \<Rightarrow> ereal" | 
| 49664 | 951 | assumes f: "\<And>i. 0 \<le> f i" | 
| 952 | and fin: "(\<Sum>i. f i) \<noteq> \<infinity>" | |
| 41980 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 953 | shows "summable (\<lambda>i. real (f i))" | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 954 | proof (rule summable_def[THEN iffD2]) | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 955 | have "0 \<le> (\<Sum>i. f i)" using assms by (auto intro: suminf_0_le) | 
| 43920 | 956 | 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 | 957 |   { 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 | 958 | 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 | 959 | note fin = this | 
| 43920 | 960 | have "(\<lambda>i. ereal (real (f i))) sums (\<Sum>i. ereal (real (f i)))" | 
| 961 | using f by (auto intro!: summable_ereal_pos summable_sums simp: ereal_le_real_iff zero_ereal_def) | |
| 962 | also have "\<dots> = ereal r" using fin r by (auto simp: ereal_real) | |
| 963 | 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 | 964 | qed | 
| 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 hoelzl parents: diff
changeset | 965 | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 966 | lemma suminf_SUP_eq: | 
| 43920 | 967 | fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ereal" | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 968 | 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 | 969 | 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 | 970 | proof - | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 971 |   { fix n :: nat
 | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 972 | have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)" | 
| 43920 | 973 | using assms by (auto intro!: SUPR_ereal_setsum[symmetric]) } | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 974 | note * = this | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 975 | show ?thesis using assms | 
| 43920 | 976 | apply (subst (1 2) suminf_ereal_eq_SUPR) | 
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 977 | unfolding * | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44918diff
changeset | 978 | apply (auto intro!: SUP_upper2) | 
| 49664 | 979 | apply (subst SUP_commute) | 
| 980 | apply rule | |
| 981 | done | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 982 | qed | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
41983diff
changeset | 983 | |
| 47761 | 984 | lemma suminf_setsum_ereal: | 
| 985 | fixes f :: "_ \<Rightarrow> _ \<Rightarrow> ereal" | |
| 986 | assumes nonneg: "\<And>i a. a \<in> A \<Longrightarrow> 0 \<le> f i a" | |
| 987 | shows "(\<Sum>i. \<Sum>a\<in>A. f i a) = (\<Sum>a\<in>A. \<Sum>i. f i a)" | |
| 988 | proof cases | |
| 49664 | 989 | assume "finite A" | 
| 990 | then show ?thesis using nonneg | |
| 47761 | 991 | by induct (simp_all add: suminf_add_ereal setsum_nonneg) | 
| 992 | qed simp | |
| 993 | ||
| 50104 | 994 | lemma suminf_ereal_eq_0: | 
| 995 | fixes f :: "nat \<Rightarrow> ereal" | |
| 996 | assumes nneg: "\<And>i. 0 \<le> f i" | |
| 997 | shows "(\<Sum>i. f i) = 0 \<longleftrightarrow> (\<forall>i. f i = 0)" | |
| 998 | proof | |
| 999 | assume "(\<Sum>i. f i) = 0" | |
| 1000 |   { fix i assume "f i \<noteq> 0"
 | |
| 1001 | with nneg have "0 < f i" by (auto simp: less_le) | |
| 1002 | also have "f i = (\<Sum>j. if j = i then f i else 0)" | |
| 1003 |       by (subst suminf_finite[where N="{i}"]) auto
 | |
| 1004 | also have "\<dots> \<le> (\<Sum>i. f i)" | |
| 1005 | using nneg by (auto intro!: suminf_le_pos) | |
| 1006 | finally have False using `(\<Sum>i. f i) = 0` by auto } | |
| 1007 | then show "\<forall>i. f i = 0" by auto | |
| 1008 | qed simp | |
| 1009 | ||
| 51340 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1010 | lemma Liminf_within: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1011 | fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1012 |   shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1013 | unfolding Liminf_def eventually_within | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1014 | proof (rule SUPR_eq, simp_all add: Ball_def Bex_def, safe) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1015 | fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1016 |   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1017 | by (auto simp: zero_less_dist_iff dist_commute) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1018 |   then show "\<exists>r>0. INFI (Collect P) f \<le> INFI (S \<inter> ball x r - {x}) f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1019 | by (intro exI[of _ d] INF_mono conjI `0 < d`) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1020 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1021 | fix d :: real assume "0 < d" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1022 | then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> P xa) \<and> | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1023 |     INFI (S \<inter> ball x d - {x}) f \<le> INFI (Collect P) f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1024 |     by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"])
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1025 | (auto intro!: INF_mono exI[of _ d] simp: dist_commute) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1026 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1027 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1028 | lemma Limsup_within: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1029 | fixes f :: "'a::metric_space => 'b::complete_lattice" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1030 |   shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1031 | unfolding Limsup_def eventually_within | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1032 | proof (rule INFI_eq, simp_all add: Ball_def Bex_def, safe) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1033 | fix P d assume "0 < d" "\<forall>y. y \<in> S \<longrightarrow> 0 < dist y x \<and> dist y x < d \<longrightarrow> P y" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1034 |   then have "S \<inter> ball x d - {x} \<subseteq> {x. P x}"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1035 | by (auto simp: zero_less_dist_iff dist_commute) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1036 |   then show "\<exists>r>0. SUPR (S \<inter> ball x r - {x}) f \<le> SUPR (Collect P) f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1037 | by (intro exI[of _ d] SUP_mono conjI `0 < d`) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1038 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1039 | fix d :: real assume "0 < d" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1040 | then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> P xa) \<and> | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1041 |     SUPR (Collect P) f \<le> SUPR (S \<inter> ball x d - {x}) f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1042 |     by (intro exI[of _ "\<lambda>y. y \<in> S \<inter> ball x d - {x}"])
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1043 | (auto intro!: SUP_mono exI[of _ d] simp: dist_commute) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1044 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1045 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1046 | lemma Liminf_at: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1047 | fixes f :: "'a::metric_space => _" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1048 |   shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1049 | using Liminf_within[of x UNIV f] by simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1050 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1051 | lemma Limsup_at: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1052 | fixes f :: "'a::metric_space => _" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1053 |   shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1054 | using Limsup_within[of x UNIV f] by simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1055 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1056 | lemma min_Liminf_at: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1057 | fixes f :: "'a::metric_space => 'b::complete_linorder" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1058 |   shows "min (f x) (Liminf (at x) f) = (SUP e:{0<..}. INF y:ball x e. f y)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1059 | unfolding inf_min[symmetric] Liminf_at | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1060 | apply (subst inf_commute) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1061 | apply (subst SUP_inf) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1062 | apply (intro SUP_cong[OF refl]) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1063 |   apply (cut_tac A="ball x b - {x}" and B="{x}" and M=f in INF_union)
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1064 | apply (simp add: INF_def del: inf_ereal_def) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1065 | done | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1066 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1067 | subsection {* monoset *}
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1068 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1069 | definition (in order) mono_set: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1070 | "mono_set S \<longleftrightarrow> (\<forall>x y. x \<le> y \<longrightarrow> x \<in> S \<longrightarrow> y \<in> S)" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1071 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1072 | lemma (in order) mono_greaterThan [intro, simp]: "mono_set {B<..}" unfolding mono_set by auto
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1073 | lemma (in order) mono_atLeast [intro, simp]: "mono_set {B..}" unfolding mono_set by auto
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1074 | lemma (in order) mono_UNIV [intro, simp]: "mono_set UNIV" unfolding mono_set by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1075 | lemma (in order) mono_empty [intro, simp]: "mono_set {}" unfolding mono_set by auto
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1076 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1077 | lemma (in complete_linorder) mono_set_iff: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1078 | fixes S :: "'a set" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1079 | defines "a \<equiv> Inf S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1080 |   shows "mono_set S \<longleftrightarrow> (S = {a <..} \<or> S = {a..})" (is "_ = ?c")
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1081 | proof | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1082 | assume "mono_set S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1083 | then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S" by (auto simp: mono_set) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1084 | show ?c | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1085 | proof cases | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1086 | assume "a \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1087 | show ?c | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1088 | using mono[OF _ `a \<in> S`] | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1089 | by (auto intro: Inf_lower simp: a_def) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1090 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1091 | assume "a \<notin> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1092 |     have "S = {a <..}"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1093 | proof safe | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1094 | fix x assume "x \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1095 | then have "a \<le> x" unfolding a_def by (rule Inf_lower) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1096 | then show "a < x" using `x \<in> S` `a \<notin> S` by (cases "a = x") auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1097 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1098 | fix x assume "a < x" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1099 | then obtain y where "y < x" "y \<in> S" unfolding a_def Inf_less_iff .. | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1100 | with mono[of y x] show "x \<in> S" by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1101 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1102 | then show ?c .. | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1103 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1104 | qed auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1105 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1106 | lemma ereal_open_mono_set: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1107 | fixes S :: "ereal set" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1108 |   shows "(open S \<and> mono_set S) \<longleftrightarrow> (S = UNIV \<or> S = {Inf S <..})"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1109 | by (metis Inf_UNIV atLeast_eq_UNIV_iff ereal_open_atLeast | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1110 | ereal_open_closed mono_set_iff open_ereal_greaterThan) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1111 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1112 | lemma ereal_closed_mono_set: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1113 | fixes S :: "ereal set" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1114 |   shows "(closed S \<and> mono_set S) \<longleftrightarrow> (S = {} \<or> S = {Inf S ..})"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1115 | by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1116 | ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1117 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1118 | lemma ereal_Liminf_Sup_monoset: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1119 | fixes f :: "'a => ereal" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1120 | shows "Liminf net f = | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1121 |     Sup {l. \<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1122 | (is "_ = Sup ?A") | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1123 | proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1124 | fix P assume P: "eventually P net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1125 | fix S assume S: "mono_set S" "INFI (Collect P) f \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1126 |   { fix x assume "P x"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1127 | then have "INFI (Collect P) f \<le> f x" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1128 | by (intro complete_lattice_class.INF_lower) simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1129 | with S have "f x \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1130 | by (simp add: mono_set) } | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1131 | with P show "eventually (\<lambda>x. f x \<in> S) net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1132 | by (auto elim: eventually_elim1) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1133 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1134 | fix y l | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1135 | assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1136 | assume P: "\<forall>P. eventually P net \<longrightarrow> INFI (Collect P) f \<le> y" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1137 | show "l \<le> y" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1138 | proof (rule dense_le) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1139 | fix B assume "B < l" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1140 |     then have "eventually (\<lambda>x. f x \<in> {B <..}) net"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1141 | by (intro S[rule_format]) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1142 |     then have "INFI {x. B < f x} f \<le> y"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1143 | using P by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1144 |     moreover have "B \<le> INFI {x. B < f x} f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1145 | by (intro INF_greatest) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1146 | ultimately show "B \<le> y" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1147 | by simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1148 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1149 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1150 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1151 | lemma ereal_Limsup_Inf_monoset: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1152 | fixes f :: "'a => ereal" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1153 | shows "Limsup net f = | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1154 |     Inf {l. \<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net}"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1155 | (is "_ = Inf ?A") | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1156 | proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1157 | fix P assume P: "eventually P net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1158 | fix S assume S: "mono_set (uminus`S)" "SUPR (Collect P) f \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1159 |   { fix x assume "P x"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1160 | then have "f x \<le> SUPR (Collect P) f" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1161 | by (intro complete_lattice_class.SUP_upper) simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1162 | with S(1)[unfolded mono_set, rule_format, of "- SUPR (Collect P) f" "- f x"] S(2) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1163 | have "f x \<in> S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1164 | by (simp add: inj_image_mem_iff) } | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1165 | with P show "eventually (\<lambda>x. f x \<in> S) net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1166 | by (auto elim: eventually_elim1) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1167 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1168 | fix y l | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1169 | assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1170 | assume P: "\<forall>P. eventually P net \<longrightarrow> y \<le> SUPR (Collect P) f" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1171 | show "y \<le> l" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1172 | proof (rule dense_ge) | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1173 | fix B assume "l < B" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1174 |     then have "eventually (\<lambda>x. f x \<in> {..< B}) net"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1175 | by (intro S[rule_format]) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1176 |     then have "y \<le> SUPR {x. f x < B} f"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1177 | using P by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1178 |     moreover have "SUPR {x. f x < B} f \<le> B"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1179 | by (intro SUP_least) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1180 | ultimately show "y \<le> B" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1181 | by simp | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1182 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1183 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1184 | |
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1185 | lemma liminf_bounded_open: | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1186 | fixes x :: "nat \<Rightarrow> ereal" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1187 | 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))" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1188 | (is "_ \<longleftrightarrow> ?P x0") | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1189 | proof | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1190 | assume "?P x0" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1191 | then show "x0 \<le> liminf x" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1192 | unfolding ereal_Liminf_Sup_monoset eventually_sequentially | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1193 | by (intro complete_lattice_class.Sup_upper) auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1194 | next | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1195 | assume "x0 \<le> liminf x" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1196 |   { fix S :: "ereal set"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1197 | assume om: "open S & mono_set S & x0:S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1198 |     { assume "S = UNIV" then have "EX N. (ALL n>=N. x n : S)" by auto }
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1199 | moreover | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1200 |     { assume "~(S=UNIV)"
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1201 |       then obtain B where B_def: "S = {B<..}" using om ereal_open_mono_set by auto
 | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1202 | then have "B<x0" using om by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1203 | then have "EX N. ALL n>=N. x n : S" | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1204 | unfolding B_def using `x0 \<le> liminf x` liminf_bounded_iff by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1205 | } | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1206 | ultimately have "EX N. (ALL n>=N. x n : S)" by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1207 | } | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1208 | then show "?P x0" by auto | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1209 | qed | 
| 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 hoelzl parents: 
51329diff
changeset | 1210 | |
| 44125 | 1211 | end |