| author | wenzelm | 
| Fri, 05 Jun 2015 11:11:26 +0200 | |
| changeset 60369 | f393a3fe884c | 
| parent 59452 | 2538b2c51769 | 
| child 60420 | 884f54e01427 | 
| 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  | 
|
| 58877 | 8  | 
section {* Limits on the Extended real number line *}
 | 
| 
41980
 
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  | 
| 
57446
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
11  | 
imports Topology_Euclidean_Space "~~/src/HOL/Library/Extended_Real" "~~/src/HOL/Library/Indicator_Function"  | 
| 
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  | 
|
| 51351 | 14  | 
lemma convergent_limsup_cl:  | 
| 53788 | 15  | 
  fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}"
 | 
| 51351 | 16  | 
shows "convergent X \<Longrightarrow> limsup X = lim X"  | 
17  | 
by (auto simp: convergent_def limI lim_imp_Limsup)  | 
|
18  | 
||
19  | 
lemma lim_increasing_cl:  | 
|
20  | 
assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<ge> f m"  | 
|
| 53788 | 21  | 
  obtains l where "f ----> (l::'a::{complete_linorder,linorder_topology})"
 | 
| 51351 | 22  | 
proof  | 
23  | 
show "f ----> (SUP n. f n)"  | 
|
24  | 
using assms  | 
|
25  | 
by (intro increasing_tendsto)  | 
|
26  | 
(auto simp: SUP_upper eventually_sequentially less_SUP_iff intro: less_le_trans)  | 
|
27  | 
qed  | 
|
28  | 
||
29  | 
lemma lim_decreasing_cl:  | 
|
30  | 
assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<le> f m"  | 
|
| 53788 | 31  | 
  obtains l where "f ----> (l::'a::{complete_linorder,linorder_topology})"
 | 
| 51351 | 32  | 
proof  | 
33  | 
show "f ----> (INF n. f n)"  | 
|
34  | 
using assms  | 
|
35  | 
by (intro decreasing_tendsto)  | 
|
36  | 
(auto simp: INF_lower eventually_sequentially INF_less_iff intro: le_less_trans)  | 
|
37  | 
qed  | 
|
38  | 
||
39  | 
lemma compact_complete_linorder:  | 
|
| 53788 | 40  | 
  fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}"
 | 
| 51351 | 41  | 
shows "\<exists>l r. subseq r \<and> (X \<circ> r) ----> l"  | 
42  | 
proof -  | 
|
43  | 
obtain r where "subseq r" and mono: "monoseq (X \<circ> r)"  | 
|
| 53788 | 44  | 
using seq_monosub[of X]  | 
45  | 
unfolding comp_def  | 
|
46  | 
by auto  | 
|
| 51351 | 47  | 
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)"  | 
48  | 
by (auto simp add: monoseq_def)  | 
|
| 53788 | 49  | 
then obtain l where "(X \<circ> r) ----> l"  | 
50  | 
using lim_increasing_cl[of "X \<circ> r"] lim_decreasing_cl[of "X \<circ> r"]  | 
|
51  | 
by auto  | 
|
52  | 
then show ?thesis  | 
|
53  | 
using `subseq r` by auto  | 
|
| 51351 | 54  | 
qed  | 
55  | 
||
| 53788 | 56  | 
lemma compact_UNIV:  | 
57  | 
  "compact (UNIV :: 'a::{complete_linorder,linorder_topology,second_countable_topology} set)"
 | 
|
| 51351 | 58  | 
using compact_complete_linorder  | 
59  | 
by (auto simp: seq_compact_eq_compact[symmetric] seq_compact_def)  | 
|
60  | 
||
61  | 
lemma compact_eq_closed:  | 
|
| 53788 | 62  | 
  fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
 | 
| 51351 | 63  | 
shows "compact S \<longleftrightarrow> closed S"  | 
| 53788 | 64  | 
using closed_inter_compact[of S, OF _ compact_UNIV] compact_imp_closed  | 
65  | 
by auto  | 
|
| 51351 | 66  | 
|
67  | 
lemma closed_contains_Sup_cl:  | 
|
| 53788 | 68  | 
  fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
 | 
69  | 
assumes "closed S"  | 
|
70  | 
    and "S \<noteq> {}"
 | 
|
71  | 
shows "Sup S \<in> S"  | 
|
| 51351 | 72  | 
proof -  | 
73  | 
from compact_eq_closed[of S] compact_attains_sup[of S] assms  | 
|
| 53788 | 74  | 
obtain s where S: "s \<in> S" "\<forall>t\<in>S. t \<le> s"  | 
75  | 
by auto  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51641 
diff
changeset
 | 
76  | 
then have "Sup S = s"  | 
| 51351 | 77  | 
by (auto intro!: Sup_eqI)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51641 
diff
changeset
 | 
78  | 
with S show ?thesis  | 
| 51351 | 79  | 
by simp  | 
80  | 
qed  | 
|
81  | 
||
82  | 
lemma closed_contains_Inf_cl:  | 
|
| 53788 | 83  | 
  fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set"
 | 
84  | 
assumes "closed S"  | 
|
85  | 
    and "S \<noteq> {}"
 | 
|
86  | 
shows "Inf S \<in> S"  | 
|
| 51351 | 87  | 
proof -  | 
88  | 
from compact_eq_closed[of S] compact_attains_inf[of S] assms  | 
|
| 53788 | 89  | 
obtain s where S: "s \<in> S" "\<forall>t\<in>S. s \<le> t"  | 
90  | 
by auto  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51641 
diff
changeset
 | 
91  | 
then have "Inf S = s"  | 
| 51351 | 92  | 
by (auto intro!: Inf_eqI)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
51641 
diff
changeset
 | 
93  | 
with S show ?thesis  | 
| 51351 | 94  | 
by simp  | 
95  | 
qed  | 
|
96  | 
||
| 53788 | 97  | 
lemma ereal_dense3:  | 
98  | 
fixes x y :: ereal  | 
|
99  | 
shows "x < y \<Longrightarrow> \<exists>r::rat. x < real_of_rat r \<and> real_of_rat r < y"  | 
|
| 51351 | 100  | 
proof (cases x y rule: ereal2_cases, simp_all)  | 
| 53788 | 101  | 
fix r q :: real  | 
102  | 
assume "r < q"  | 
|
103  | 
from Rats_dense_in_real[OF this] show "\<exists>x. r < real_of_rat x \<and> real_of_rat x < q"  | 
|
| 51351 | 104  | 
by (fastforce simp: Rats_def)  | 
105  | 
next  | 
|
| 53788 | 106  | 
fix r :: real  | 
107  | 
show "\<exists>x. r < real_of_rat x" "\<exists>x. real_of_rat x < r"  | 
|
| 51351 | 108  | 
using gt_ex[of r] lt_ex[of r] Rats_dense_in_real  | 
109  | 
by (auto simp: Rats_def)  | 
|
110  | 
qed  | 
|
111  | 
||
112  | 
instance ereal :: second_countable_topology  | 
|
113  | 
proof (default, intro exI conjI)  | 
|
114  | 
  let ?B = "(\<Union>r\<in>\<rat>. {{..< r}, {r <..}} :: ereal set set)"
 | 
|
| 53788 | 115  | 
show "countable ?B"  | 
116  | 
by (auto intro: countable_rat)  | 
|
| 51351 | 117  | 
show "open = generate_topology ?B"  | 
118  | 
proof (intro ext iffI)  | 
|
| 53788 | 119  | 
fix S :: "ereal set"  | 
120  | 
assume "open S"  | 
|
| 51351 | 121  | 
then show "generate_topology ?B S"  | 
122  | 
unfolding open_generated_order  | 
|
123  | 
proof induct  | 
|
124  | 
case (Basis b)  | 
|
| 53788 | 125  | 
      then obtain e where "b = {..<e} \<or> b = {e<..}"
 | 
126  | 
by auto  | 
|
| 51351 | 127  | 
      moreover have "{..<e} = \<Union>{{..<x}|x. x \<in> \<rat> \<and> x < e}" "{e<..} = \<Union>{{x<..}|x. x \<in> \<rat> \<and> e < x}"
 | 
128  | 
by (auto dest: ereal_dense3  | 
|
129  | 
simp del: ex_simps  | 
|
130  | 
simp add: ex_simps[symmetric] conj_commute Rats_def image_iff)  | 
|
131  | 
ultimately show ?case  | 
|
132  | 
by (auto intro: generate_topology.intros)  | 
|
133  | 
qed (auto intro: generate_topology.intros)  | 
|
134  | 
next  | 
|
| 53788 | 135  | 
fix S  | 
136  | 
assume "generate_topology ?B S"  | 
|
137  | 
then show "open S"  | 
|
138  | 
by induct auto  | 
|
| 51351 | 139  | 
qed  | 
140  | 
qed  | 
|
141  | 
||
| 43920 | 142  | 
lemma continuous_on_ereal[intro, simp]: "continuous_on A ereal"  | 
| 53788 | 143  | 
unfolding continuous_on_topological open_ereal_def  | 
144  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
145  | 
|
| 43920 | 146  | 
lemma continuous_at_ereal[intro, simp]: "continuous (at x) ereal"  | 
| 53788 | 147  | 
using continuous_on_eq_continuous_at[of UNIV]  | 
148  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
149  | 
|
| 43920 | 150  | 
lemma continuous_within_ereal[intro, simp]: "x \<in> A \<Longrightarrow> continuous (at x within A) ereal"  | 
| 53788 | 151  | 
using continuous_on_eq_continuous_within[of A]  | 
152  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
153  | 
|
| 43920 | 154  | 
lemma ereal_open_uminus:  | 
155  | 
fixes S :: "ereal set"  | 
|
| 53788 | 156  | 
assumes "open S"  | 
157  | 
shows "open (uminus ` S)"  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
158  | 
using `open S`[unfolded open_generated_order]  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
159  | 
proof induct  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
160  | 
have "range uminus = (UNIV :: ereal set)"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
161  | 
by (auto simp: image_iff ereal_uminus_eq_reorder)  | 
| 53788 | 162  | 
then show "open (range uminus :: ereal set)"  | 
163  | 
by simp  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
164  | 
qed (auto simp add: image_Union image_Int)  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
165  | 
|
| 43920 | 166  | 
lemma ereal_uminus_complement:  | 
167  | 
fixes S :: "ereal set"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
168  | 
shows "uminus ` (- S) = - uminus ` S"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
169  | 
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
 | 
170  | 
|
| 43920 | 171  | 
lemma ereal_closed_uminus:  | 
172  | 
fixes S :: "ereal set"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
173  | 
assumes "closed S"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
174  | 
shows "closed (uminus ` S)"  | 
| 53788 | 175  | 
using assms  | 
176  | 
unfolding closed_def ereal_uminus_complement[symmetric]  | 
|
177  | 
by (rule ereal_open_uminus)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
178  | 
|
| 43920 | 179  | 
lemma ereal_open_closed_aux:  | 
180  | 
fixes S :: "ereal set"  | 
|
| 53788 | 181  | 
assumes "open S"  | 
182  | 
and "closed S"  | 
|
183  | 
and S: "(-\<infinity>) \<notin> S"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
184  | 
  shows "S = {}"
 | 
| 49664 | 185  | 
proof (rule ccontr)  | 
| 53788 | 186  | 
assume "\<not> ?thesis"  | 
187  | 
then have *: "Inf S \<in> S"  | 
|
188  | 
by (metis assms(2) closed_contains_Inf_cl)  | 
|
189  | 
  {
 | 
|
190  | 
assume "Inf S = -\<infinity>"  | 
|
191  | 
then have False  | 
|
192  | 
using * assms(3) by auto  | 
|
193  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
194  | 
moreover  | 
| 53788 | 195  | 
  {
 | 
196  | 
assume "Inf S = \<infinity>"  | 
|
197  | 
    then have "S = {\<infinity>}"
 | 
|
198  | 
      by (metis Inf_eq_PInfty `S \<noteq> {}`)
 | 
|
199  | 
then have False  | 
|
200  | 
by (metis assms(1) not_open_singleton)  | 
|
201  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
202  | 
moreover  | 
| 53788 | 203  | 
  {
 | 
204  | 
assume fin: "\<bar>Inf S\<bar> \<noteq> \<infinity>"  | 
|
205  | 
from ereal_open_cont_interval[OF assms(1) * fin]  | 
|
206  | 
    obtain e where e: "e > 0" "{Inf S - e<..<Inf S + e} \<subseteq> S" .
 | 
|
207  | 
then obtain b where b: "Inf S - e < b" "b < Inf S"  | 
|
208  | 
using fin ereal_between[of "Inf S" e] dense[of "Inf S - e"]  | 
|
| 44918 | 209  | 
by auto  | 
| 53788 | 210  | 
    then have "b: {Inf S - e <..< Inf S + e}"
 | 
211  | 
using e fin ereal_between[of "Inf S" e]  | 
|
212  | 
by auto  | 
|
213  | 
then have "b \<in> S"  | 
|
214  | 
using e by auto  | 
|
215  | 
then have False  | 
|
216  | 
using b by (metis complete_lattice_class.Inf_lower leD)  | 
|
217  | 
}  | 
|
218  | 
ultimately show False  | 
|
219  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
220  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
221  | 
|
| 43920 | 222  | 
lemma ereal_open_closed:  | 
223  | 
fixes S :: "ereal set"  | 
|
| 53788 | 224  | 
  shows "open S \<and> closed S \<longleftrightarrow> S = {} \<or> S = UNIV"
 | 
| 49664 | 225  | 
proof -  | 
| 53788 | 226  | 
  {
 | 
227  | 
assume lhs: "open S \<and> closed S"  | 
|
228  | 
    {
 | 
|
229  | 
assume "-\<infinity> \<notin> S"  | 
|
230  | 
      then have "S = {}"
 | 
|
231  | 
using lhs ereal_open_closed_aux by auto  | 
|
232  | 
}  | 
|
| 49664 | 233  | 
moreover  | 
| 53788 | 234  | 
    {
 | 
235  | 
assume "-\<infinity> \<in> S"  | 
|
236  | 
      then have "- S = {}"
 | 
|
237  | 
using lhs ereal_open_closed_aux[of "-S"] by auto  | 
|
238  | 
}  | 
|
239  | 
    ultimately have "S = {} \<or> S = UNIV"
 | 
|
240  | 
by auto  | 
|
241  | 
}  | 
|
242  | 
then show ?thesis  | 
|
243  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
244  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
245  | 
|
| 43920 | 246  | 
lemma ereal_open_affinity_pos:  | 
| 43923 | 247  | 
fixes S :: "ereal set"  | 
| 53788 | 248  | 
assumes "open S"  | 
249  | 
and m: "m \<noteq> \<infinity>" "0 < m"  | 
|
250  | 
and t: "\<bar>t\<bar> \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
251  | 
shows "open ((\<lambda>x. m * x + t) ` S)"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
252  | 
proof -  | 
| 
59452
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
253  | 
have "open ((\<lambda>x. inverse m * (x + -t)) -` S)"  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
254  | 
using m t  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
255  | 
apply (intro open_vimage `open S`)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
256  | 
apply (intro continuous_at_imp_continuous_on ballI tendsto_cmult_ereal continuous_at[THEN iffD2]  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
257  | 
tendsto_ident_at tendsto_add_left_ereal)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
258  | 
apply auto  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
259  | 
done  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
260  | 
also have "(\<lambda>x. inverse m * (x + -t)) -` S = (\<lambda>x. (x - t) / m) -` S"  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
261  | 
using m t by (auto simp: divide_ereal_def mult.commute uminus_ereal.simps[symmetric] minus_ereal_def  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
262  | 
simp del: uminus_ereal.simps)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
263  | 
also have "(\<lambda>x. (x - t) / m) -` S = (\<lambda>x. m * x + t) ` S"  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
264  | 
using m t  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
265  | 
by (simp add: set_eq_iff image_iff)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
266  | 
(metis abs_ereal_less0 abs_ereal_uminus ereal_divide_eq ereal_eq_minus ereal_minus(7,8)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
267  | 
ereal_minus_less_minus ereal_mult_eq_PInfty ereal_uminus_uminus ereal_zero_mult)  | 
| 
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
268  | 
finally show ?thesis .  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
269  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
270  | 
|
| 43920 | 271  | 
lemma ereal_open_affinity:  | 
| 43923 | 272  | 
fixes S :: "ereal set"  | 
| 49664 | 273  | 
assumes "open S"  | 
274  | 
and m: "\<bar>m\<bar> \<noteq> \<infinity>" "m \<noteq> 0"  | 
|
275  | 
and t: "\<bar>t\<bar> \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
276  | 
shows "open ((\<lambda>x. m * x + t) ` S)"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
277  | 
proof cases  | 
| 49664 | 278  | 
assume "0 < m"  | 
279  | 
then show ?thesis  | 
|
| 53788 | 280  | 
using ereal_open_affinity_pos[OF `open S` _ _ t, of m] m  | 
281  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
282  | 
next  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
283  | 
assume "\<not> 0 < m" then  | 
| 53788 | 284  | 
have "0 < -m"  | 
285  | 
using `m \<noteq> 0`  | 
|
286  | 
by (cases m) auto  | 
|
287  | 
then have m: "-m \<noteq> \<infinity>" "0 < -m"  | 
|
288  | 
using `\<bar>m\<bar> \<noteq> \<infinity>`  | 
|
| 43920 | 289  | 
by (auto simp: ereal_uminus_eq_reorder)  | 
| 53788 | 290  | 
from ereal_open_affinity_pos[OF ereal_open_uminus[OF `open S`] m t] show ?thesis  | 
291  | 
unfolding image_image by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
292  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
293  | 
|
| 53788 | 294  | 
lemma ereal_open_atLeast:  | 
295  | 
fixes x :: ereal  | 
|
296  | 
  shows "open {x..} \<longleftrightarrow> x = -\<infinity>"
 | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
297  | 
proof  | 
| 53788 | 298  | 
assume "x = -\<infinity>"  | 
299  | 
  then have "{x..} = UNIV"
 | 
|
300  | 
by auto  | 
|
301  | 
  then show "open {x..}"
 | 
|
302  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
303  | 
next  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
304  | 
  assume "open {x..}"
 | 
| 53788 | 305  | 
  then have "open {x..} \<and> closed {x..}"
 | 
306  | 
by auto  | 
|
307  | 
  then have "{x..} = UNIV"
 | 
|
308  | 
unfolding ereal_open_closed by auto  | 
|
309  | 
then show "x = -\<infinity>"  | 
|
310  | 
by (simp add: bot_ereal_def atLeast_eq_UNIV_iff)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
311  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
312  | 
|
| 53788 | 313  | 
lemma open_uminus_iff:  | 
314  | 
fixes S :: "ereal set"  | 
|
315  | 
shows "open (uminus ` S) \<longleftrightarrow> open S"  | 
|
316  | 
using ereal_open_uminus[of S] ereal_open_uminus[of "uminus ` S"]  | 
|
317  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
318  | 
|
| 43920 | 319  | 
lemma ereal_Liminf_uminus:  | 
| 53788 | 320  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
321  | 
shows "Liminf net (\<lambda>x. - (f x)) = - Limsup net f"  | 
|
| 43920 | 322  | 
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
 | 
323  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
324  | 
lemma Liminf_PInfty:  | 
| 43920 | 325  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
326  | 
assumes "\<not> trivial_limit net"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
327  | 
shows "(f ---> \<infinity>) net \<longleftrightarrow> Liminf net f = \<infinity>"  | 
| 53788 | 328  | 
unfolding tendsto_iff_Liminf_eq_Limsup[OF assms]  | 
329  | 
using Liminf_le_Limsup[OF assms, of f]  | 
|
330  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
331  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
332  | 
lemma Limsup_MInfty:  | 
| 43920 | 333  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
334  | 
assumes "\<not> trivial_limit net"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
335  | 
shows "(f ---> -\<infinity>) net \<longleftrightarrow> Limsup net f = -\<infinity>"  | 
| 53788 | 336  | 
unfolding tendsto_iff_Liminf_eq_Limsup[OF assms]  | 
337  | 
using Liminf_le_Limsup[OF assms, of f]  | 
|
338  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
339  | 
|
| 50104 | 340  | 
lemma convergent_ereal:  | 
| 53788 | 341  | 
  fixes X :: "nat \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
 | 
| 50104 | 342  | 
shows "convergent X \<longleftrightarrow> limsup X = liminf X"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
343  | 
using tendsto_iff_Liminf_eq_Limsup[of sequentially]  | 
| 50104 | 344  | 
by (auto simp: convergent_def)  | 
345  | 
||
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
346  | 
lemma limsup_le_liminf_real:  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
347  | 
fixes X :: "nat \<Rightarrow> real" and L :: real  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
348  | 
assumes 1: "limsup X \<le> L" and 2: "L \<le> liminf X"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
349  | 
shows "X ----> L"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
350  | 
proof -  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
351  | 
from 1 2 have "limsup X \<le> liminf X" by auto  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
352  | 
hence 3: "limsup X = liminf X"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
353  | 
apply (subst eq_iff, rule conjI)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
354  | 
by (rule Liminf_le_Limsup, auto)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
355  | 
hence 4: "convergent (\<lambda>n. ereal (X n))"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
356  | 
by (subst convergent_ereal)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
357  | 
hence "limsup X = lim (\<lambda>n. ereal(X n))"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
358  | 
by (rule convergent_limsup_cl)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
359  | 
also from 1 2 3 have "limsup X = L" by auto  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
360  | 
finally have "lim (\<lambda>n. ereal(X n)) = L" ..  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
361  | 
hence "(\<lambda>n. ereal (X n)) ----> L"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
362  | 
apply (elim subst)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
363  | 
by (subst convergent_LIMSEQ_iff [symmetric], rule 4)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
364  | 
thus ?thesis by simp  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
365  | 
qed  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
366  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
367  | 
lemma liminf_PInfty:  | 
| 51351 | 368  | 
fixes X :: "nat \<Rightarrow> ereal"  | 
369  | 
shows "X ----> \<infinity> \<longleftrightarrow> liminf X = \<infinity>"  | 
|
| 49664 | 370  | 
by (metis Liminf_PInfty trivial_limit_sequentially)  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
371  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
372  | 
lemma limsup_MInfty:  | 
| 51351 | 373  | 
fixes X :: "nat \<Rightarrow> ereal"  | 
374  | 
shows "X ----> -\<infinity> \<longleftrightarrow> limsup X = -\<infinity>"  | 
|
| 49664 | 375  | 
by (metis Limsup_MInfty trivial_limit_sequentially)  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
376  | 
|
| 43920 | 377  | 
lemma ereal_lim_mono:  | 
| 53788 | 378  | 
fixes X Y :: "nat \<Rightarrow> 'a::linorder_topology"  | 
379  | 
assumes "\<And>n. N \<le> n \<Longrightarrow> X n \<le> Y n"  | 
|
380  | 
and "X ----> x"  | 
|
381  | 
and "Y ----> y"  | 
|
382  | 
shows "x \<le> y"  | 
|
| 51000 | 383  | 
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
 | 
384  | 
|
| 43920 | 385  | 
lemma incseq_le_ereal:  | 
| 51351 | 386  | 
fixes X :: "nat \<Rightarrow> 'a::linorder_topology"  | 
| 53788 | 387  | 
assumes inc: "incseq X"  | 
388  | 
and lim: "X ----> L"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
389  | 
shows "X N \<le> L"  | 
| 53788 | 390  | 
using inc  | 
391  | 
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
 | 
392  | 
|
| 49664 | 393  | 
lemma decseq_ge_ereal:  | 
394  | 
assumes dec: "decseq X"  | 
|
| 51351 | 395  | 
and lim: "X ----> (L::'a::linorder_topology)"  | 
| 53788 | 396  | 
shows "X N \<ge> L"  | 
| 49664 | 397  | 
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
 | 
398  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
399  | 
lemma bounded_abs:  | 
| 53788 | 400  | 
fixes a :: real  | 
401  | 
assumes "a \<le> x"  | 
|
402  | 
and "x \<le> b"  | 
|
403  | 
shows "abs x \<le> max (abs a) (abs b)"  | 
|
| 49664 | 404  | 
by (metis abs_less_iff assms leI le_max_iff_disj  | 
405  | 
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
 | 
406  | 
|
| 43920 | 407  | 
lemma ereal_Sup_lim:  | 
| 53788 | 408  | 
  fixes a :: "'a::{complete_linorder,linorder_topology}"
 | 
409  | 
assumes "\<And>n. b n \<in> s"  | 
|
410  | 
and "b ----> a"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
411  | 
shows "a \<le> Sup s"  | 
| 49664 | 412  | 
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
 | 
413  | 
|
| 43920 | 414  | 
lemma ereal_Inf_lim:  | 
| 53788 | 415  | 
  fixes a :: "'a::{complete_linorder,linorder_topology}"
 | 
416  | 
assumes "\<And>n. b n \<in> s"  | 
|
417  | 
and "b ----> a"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
418  | 
shows "Inf s \<le> a"  | 
| 49664 | 419  | 
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
 | 
420  | 
|
| 43920 | 421  | 
lemma SUP_Lim_ereal:  | 
| 53788 | 422  | 
  fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}"
 | 
423  | 
assumes inc: "incseq X"  | 
|
424  | 
and l: "X ----> l"  | 
|
425  | 
shows "(SUP n. X n) = l"  | 
|
426  | 
using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l]  | 
|
427  | 
by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
428  | 
|
| 51351 | 429  | 
lemma INF_Lim_ereal:  | 
| 53788 | 430  | 
  fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}"
 | 
431  | 
assumes dec: "decseq X"  | 
|
432  | 
and l: "X ----> l"  | 
|
433  | 
shows "(INF n. X n) = l"  | 
|
434  | 
using LIMSEQ_INF[OF dec] tendsto_unique[OF trivial_limit_sequentially l]  | 
|
435  | 
by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
436  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
437  | 
lemma SUP_eq_LIMSEQ:  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
438  | 
assumes "mono f"  | 
| 43920 | 439  | 
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
 | 
440  | 
proof  | 
| 43920 | 441  | 
have inc: "incseq (\<lambda>i. ereal (f i))"  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
442  | 
using `mono f` unfolding mono_def incseq_def by auto  | 
| 53788 | 443  | 
  {
 | 
444  | 
assume "f ----> x"  | 
|
445  | 
then have "(\<lambda>i. ereal (f i)) ----> ereal x"  | 
|
446  | 
by auto  | 
|
447  | 
from SUP_Lim_ereal[OF inc this] show "(SUP n. ereal (f n)) = ereal x" .  | 
|
448  | 
next  | 
|
449  | 
assume "(SUP n. ereal (f n)) = ereal x"  | 
|
450  | 
with LIMSEQ_SUP[OF inc] show "f ----> x" by auto  | 
|
451  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
452  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
453  | 
|
| 43920 | 454  | 
lemma liminf_ereal_cminus:  | 
| 49664 | 455  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
456  | 
assumes "c \<noteq> -\<infinity>"  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
457  | 
shows "liminf (\<lambda>x. c - f x) = c - limsup f"  | 
| 
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
458  | 
proof (cases c)  | 
| 49664 | 459  | 
case PInf  | 
| 53788 | 460  | 
then show ?thesis  | 
461  | 
by (simp add: Liminf_const)  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
462  | 
next  | 
| 49664 | 463  | 
case (real r)  | 
464  | 
then show ?thesis  | 
|
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
465  | 
unfolding liminf_SUP_INF limsup_INF_SUP  | 
| 
59452
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
466  | 
apply (subst INF_ereal_minus_right)  | 
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
467  | 
apply auto  | 
| 
59452
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
468  | 
apply (subst SUP_ereal_minus_right)  | 
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
469  | 
apply auto  | 
| 
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
470  | 
done  | 
| 
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
471  | 
qed (insert `c \<noteq> -\<infinity>`, simp)  | 
| 
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
472  | 
|
| 49664 | 473  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
474  | 
subsubsection {* Continuity *}
 | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
475  | 
|
| 43920 | 476  | 
lemma continuous_at_of_ereal:  | 
477  | 
fixes x0 :: ereal  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
478  | 
assumes "\<bar>x0\<bar> \<noteq> \<infinity>"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
479  | 
shows "continuous (at x0) real"  | 
| 49664 | 480  | 
proof -  | 
| 53788 | 481  | 
  {
 | 
482  | 
fix T  | 
|
483  | 
assume T: "open T" "real x0 \<in> T"  | 
|
484  | 
def S \<equiv> "ereal ` T"  | 
|
485  | 
then have "ereal (real x0) \<in> S"  | 
|
486  | 
using T by auto  | 
|
487  | 
then have "x0 \<in> S"  | 
|
488  | 
using assms ereal_real by auto  | 
|
489  | 
moreover have "open S"  | 
|
490  | 
using open_ereal S_def T by auto  | 
|
491  | 
moreover have "\<forall>y\<in>S. real y \<in> T"  | 
|
492  | 
using S_def T by auto  | 
|
493  | 
ultimately have "\<exists>S. x0 \<in> S \<and> open S \<and> (\<forall>y\<in>S. real y \<in> T)"  | 
|
494  | 
by auto  | 
|
| 49664 | 495  | 
}  | 
| 53788 | 496  | 
then show ?thesis  | 
497  | 
unfolding continuous_at_open by blast  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
498  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
499  | 
|
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
500  | 
lemma nhds_ereal: "nhds (ereal r) = filtermap ereal (nhds r)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
501  | 
by (simp add: filtermap_nhds_open_map open_ereal continuous_at_of_ereal)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
502  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
503  | 
lemma at_ereal: "at (ereal r) = filtermap ereal (at r)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
504  | 
by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
505  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
506  | 
lemma at_left_ereal: "at_left (ereal r) = filtermap ereal (at_left r)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
507  | 
by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
508  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
509  | 
lemma at_right_ereal: "at_right (ereal r) = filtermap ereal (at_right r)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
510  | 
by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
511  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
512  | 
lemma  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
513  | 
shows at_left_PInf: "at_left \<infinity> = filtermap ereal at_top"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
514  | 
and at_right_MInf: "at_right (-\<infinity>) = filtermap ereal at_bot"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
515  | 
unfolding filter_eq_iff eventually_filtermap eventually_at_top_dense eventually_at_bot_dense  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
516  | 
eventually_at_left[OF ereal_less(5)] eventually_at_right[OF ereal_less(6)]  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
517  | 
by (auto simp add: ereal_all_split ereal_ex_split)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
518  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
519  | 
lemma ereal_tendsto_simps1:  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
520  | 
"((f \<circ> real) ---> y) (at_left (ereal x)) \<longleftrightarrow> (f ---> y) (at_left x)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
521  | 
"((f \<circ> real) ---> y) (at_right (ereal x)) \<longleftrightarrow> (f ---> y) (at_right x)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
522  | 
"((f \<circ> real) ---> y) (at_left (\<infinity>::ereal)) \<longleftrightarrow> (f ---> y) at_top"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
523  | 
"((f \<circ> real) ---> y) (at_right (-\<infinity>::ereal)) \<longleftrightarrow> (f ---> y) at_bot"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
524  | 
unfolding tendsto_compose_filtermap at_left_ereal at_right_ereal at_left_PInf at_right_MInf  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
525  | 
by (auto simp: filtermap_filtermap filtermap_ident)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
526  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
527  | 
lemma ereal_tendsto_simps2:  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
528  | 
"((ereal \<circ> f) ---> ereal a) F \<longleftrightarrow> (f ---> a) F"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
529  | 
"((ereal \<circ> f) ---> \<infinity>) F \<longleftrightarrow> (LIM x F. f x :> at_top)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
530  | 
"((ereal \<circ> f) ---> -\<infinity>) F \<longleftrightarrow> (LIM x F. f x :> at_bot)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
531  | 
unfolding tendsto_PInfty filterlim_at_top_dense tendsto_MInfty filterlim_at_bot_dense  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
532  | 
using lim_ereal by (simp_all add: comp_def)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
533  | 
|
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
534  | 
lemmas ereal_tendsto_simps = ereal_tendsto_simps1 ereal_tendsto_simps2  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
535  | 
|
| 43920 | 536  | 
lemma continuous_at_iff_ereal:  | 
| 53788 | 537  | 
fixes f :: "'a::t2_space \<Rightarrow> real"  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
538  | 
shows "continuous (at x0 within s) f \<longleftrightarrow> continuous (at x0 within s) (ereal \<circ> f)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
539  | 
unfolding continuous_within comp_def lim_ereal ..  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
540  | 
|
| 43920 | 541  | 
lemma continuous_on_iff_ereal:  | 
| 49664 | 542  | 
fixes f :: "'a::t2_space => real"  | 
| 53788 | 543  | 
assumes "open A"  | 
544  | 
shows "continuous_on A f \<longleftrightarrow> continuous_on A (ereal \<circ> f)"  | 
|
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
545  | 
unfolding continuous_on_def comp_def lim_ereal ..  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
546  | 
|
| 53788 | 547  | 
lemma continuous_on_real: "continuous_on (UNIV - {\<infinity>, -\<infinity>::ereal}) real"
 | 
548  | 
using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal  | 
|
549  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
551  | 
lemma continuous_on_iff_real:  | 
| 53788 | 552  | 
fixes f :: "'a::t2_space \<Rightarrow> ereal"  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57446 
diff
changeset
 | 
553  | 
assumes *: "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
554  | 
shows "continuous_on A f \<longleftrightarrow> continuous_on A (real \<circ> f)"  | 
| 49664 | 555  | 
proof -  | 
| 53788 | 556  | 
  have "f ` A \<subseteq> UNIV - {\<infinity>, -\<infinity>}"
 | 
557  | 
using assms by force  | 
|
| 49664 | 558  | 
then have *: "continuous_on (f ` A) real"  | 
559  | 
using continuous_on_real by (simp add: continuous_on_subset)  | 
|
| 53788 | 560  | 
have **: "continuous_on ((real \<circ> f) ` A) ereal"  | 
561  | 
using continuous_on_ereal continuous_on_subset[of "UNIV" "ereal" "(real \<circ> f) ` A"]  | 
|
562  | 
by blast  | 
|
563  | 
  {
 | 
|
564  | 
assume "continuous_on A f"  | 
|
565  | 
then have "continuous_on A (real \<circ> f)"  | 
|
| 49664 | 566  | 
apply (subst continuous_on_compose)  | 
| 53788 | 567  | 
using *  | 
568  | 
apply auto  | 
|
| 49664 | 569  | 
done  | 
570  | 
}  | 
|
571  | 
moreover  | 
|
| 53788 | 572  | 
  {
 | 
573  | 
assume "continuous_on A (real \<circ> f)"  | 
|
574  | 
then have "continuous_on A (ereal \<circ> (real \<circ> f))"  | 
|
| 49664 | 575  | 
apply (subst continuous_on_compose)  | 
| 53788 | 576  | 
using **  | 
577  | 
apply auto  | 
|
| 49664 | 578  | 
done  | 
579  | 
then have "continuous_on A f"  | 
|
| 53788 | 580  | 
apply (subst continuous_on_eq[of A "ereal \<circ> (real \<circ> f)" f])  | 
581  | 
using assms ereal_real  | 
|
582  | 
apply auto  | 
|
| 49664 | 583  | 
done  | 
584  | 
}  | 
|
| 53788 | 585  | 
ultimately show ?thesis  | 
586  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
587  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
588  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
589  | 
lemma continuous_at_const:  | 
| 53788 | 590  | 
fixes f :: "'a::t2_space \<Rightarrow> ereal"  | 
591  | 
assumes "\<forall>x. f x = C"  | 
|
592  | 
shows "\<forall>x. continuous (at x) f"  | 
|
593  | 
unfolding continuous_at_open  | 
|
594  | 
using assms t1_space  | 
|
595  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
596  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
597  | 
lemma mono_closed_real:  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
598  | 
fixes S :: "real set"  | 
| 53788 | 599  | 
assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S"  | 
| 49664 | 600  | 
and "closed S"  | 
| 53788 | 601  | 
  shows "S = {} \<or> S = UNIV \<or> (\<exists>a. S = {a..})"
 | 
| 49664 | 602  | 
proof -  | 
| 53788 | 603  | 
  {
 | 
604  | 
    assume "S \<noteq> {}"
 | 
|
605  | 
    { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x"
 | 
|
606  | 
then have *: "\<forall>x\<in>S. Inf S \<le> x"  | 
|
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54257 
diff
changeset
 | 
607  | 
using cInf_lower[of _ S] ex by (metis bdd_below_def)  | 
| 53788 | 608  | 
then have "Inf S \<in> S"  | 
609  | 
apply (subst closed_contains_Inf)  | 
|
610  | 
        using ex `S \<noteq> {}` `closed S`
 | 
|
611  | 
apply auto  | 
|
612  | 
done  | 
|
613  | 
then have "\<forall>x. Inf S \<le> x \<longleftrightarrow> x \<in> S"  | 
|
614  | 
using mono[rule_format, of "Inf S"] *  | 
|
615  | 
by auto  | 
|
616  | 
      then have "S = {Inf S ..}"
 | 
|
617  | 
by auto  | 
|
618  | 
      then have "\<exists>a. S = {a ..}"
 | 
|
619  | 
by auto  | 
|
| 49664 | 620  | 
}  | 
621  | 
moreover  | 
|
| 53788 | 622  | 
    {
 | 
623  | 
assume "\<not> (\<exists>B. \<forall>x\<in>S. B \<le> x)"  | 
|
624  | 
then have nex: "\<forall>B. \<exists>x\<in>S. x < B"  | 
|
625  | 
by (simp add: not_le)  | 
|
626  | 
      {
 | 
|
627  | 
fix y  | 
|
628  | 
obtain x where "x\<in>S" and "x < y"  | 
|
629  | 
using nex by auto  | 
|
630  | 
then have "y \<in> S"  | 
|
631  | 
using mono[rule_format, of x y] by auto  | 
|
632  | 
}  | 
|
633  | 
then have "S = UNIV"  | 
|
634  | 
by auto  | 
|
| 49664 | 635  | 
}  | 
| 53788 | 636  | 
    ultimately have "S = UNIV \<or> (\<exists>a. S = {a ..})"
 | 
637  | 
by blast  | 
|
638  | 
}  | 
|
639  | 
then show ?thesis  | 
|
640  | 
by blast  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
641  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
642  | 
|
| 43920 | 643  | 
lemma mono_closed_ereal:  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
644  | 
fixes S :: "real set"  | 
| 53788 | 645  | 
assumes mono: "\<forall>y z. y \<in> S \<and> y \<le> z \<longrightarrow> z \<in> S"  | 
| 49664 | 646  | 
and "closed S"  | 
| 53788 | 647  | 
  shows "\<exists>a. S = {x. a \<le> ereal x}"
 | 
| 49664 | 648  | 
proof -  | 
| 53788 | 649  | 
  {
 | 
650  | 
    assume "S = {}"
 | 
|
651  | 
then have ?thesis  | 
|
652  | 
apply (rule_tac x=PInfty in exI)  | 
|
653  | 
apply auto  | 
|
654  | 
done  | 
|
655  | 
}  | 
|
| 49664 | 656  | 
moreover  | 
| 53788 | 657  | 
  {
 | 
658  | 
assume "S = UNIV"  | 
|
659  | 
then have ?thesis  | 
|
660  | 
apply (rule_tac x="-\<infinity>" in exI)  | 
|
661  | 
apply auto  | 
|
662  | 
done  | 
|
663  | 
}  | 
|
| 49664 | 664  | 
moreover  | 
| 53788 | 665  | 
  {
 | 
666  | 
    assume "\<exists>a. S = {a ..}"
 | 
|
667  | 
    then obtain a where "S = {a ..}"
 | 
|
668  | 
by auto  | 
|
669  | 
then have ?thesis  | 
|
670  | 
apply (rule_tac x="ereal a" in exI)  | 
|
671  | 
apply auto  | 
|
672  | 
done  | 
|
| 49664 | 673  | 
}  | 
| 53788 | 674  | 
ultimately show ?thesis  | 
675  | 
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
 | 
676  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
677  | 
|
| 53788 | 678  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
679  | 
subsection {* Sums *}
 | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
680  | 
|
| 43920 | 681  | 
lemma sums_ereal_positive:  | 
| 49664 | 682  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
683  | 
assumes "\<And>i. 0 \<le> f i"  | 
|
684  | 
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
 | 
685  | 
proof -  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
686  | 
have "incseq (\<lambda>i. \<Sum>j=0..<i. f j)"  | 
| 53788 | 687  | 
using ereal_add_mono[OF _ assms]  | 
688  | 
by (auto intro!: incseq_SucI)  | 
|
| 51000 | 689  | 
from LIMSEQ_SUP[OF this]  | 
| 53788 | 690  | 
show ?thesis unfolding sums_def  | 
691  | 
by (simp add: atLeast0LessThan)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
692  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
693  | 
|
| 43920 | 694  | 
lemma summable_ereal_pos:  | 
| 49664 | 695  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
696  | 
assumes "\<And>i. 0 \<le> f i"  | 
|
697  | 
shows "summable f"  | 
|
| 53788 | 698  | 
using sums_ereal_positive[of f, OF assms]  | 
699  | 
unfolding summable_def  | 
|
700  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
701  | 
|
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
702  | 
lemma suminf_ereal_eq_SUP:  | 
| 49664 | 703  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
704  | 
assumes "\<And>i. 0 \<le> f i"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
705  | 
shows "(\<Sum>x. f x) = (SUP n. \<Sum>i<n. f i)"  | 
| 53788 | 706  | 
using sums_ereal_positive[of f, OF assms, THEN sums_unique]  | 
707  | 
by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
708  | 
|
| 49664 | 709  | 
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
 | 
710  | 
unfolding sums_def by simp  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
711  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
712  | 
lemma suminf_bound:  | 
| 43920 | 713  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
| 53788 | 714  | 
assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x"  | 
715  | 
and pos: "\<And>n. 0 \<le> f n"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
716  | 
shows "suminf f \<le> x"  | 
| 43920 | 717  | 
proof (rule Lim_bounded_ereal)  | 
718  | 
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
 | 
719  | 
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
 | 
720  | 
by (auto dest!: summable_sums simp: sums_def atLeast0LessThan)  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
721  | 
  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
 | 
722  | 
using assms by auto  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
723  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
724  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
725  | 
lemma suminf_bound_add:  | 
| 43920 | 726  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
| 49664 | 727  | 
assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x"  | 
728  | 
and pos: "\<And>n. 0 \<le> f n"  | 
|
729  | 
and "y \<noteq> -\<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
730  | 
shows "suminf f + y \<le> x"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
731  | 
proof (cases y)  | 
| 49664 | 732  | 
case (real r)  | 
733  | 
then have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y"  | 
|
| 43920 | 734  | 
using assms by (simp add: ereal_le_minus)  | 
| 53788 | 735  | 
then have "(\<Sum> n. f n) \<le> x - y"  | 
736  | 
using pos by (rule suminf_bound)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
737  | 
then show "(\<Sum> n. f n) + y \<le> x"  | 
| 43920 | 738  | 
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
 | 
739  | 
qed (insert assms, auto)  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
740  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
741  | 
lemma suminf_upper:  | 
| 49664 | 742  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
743  | 
assumes "\<And>n. 0 \<le> f n"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
744  | 
shows "(\<Sum>n<N. f n) \<le> (\<Sum>n. f n)"  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
745  | 
unfolding suminf_ereal_eq_SUP [OF assms]  | 
| 56166 | 746  | 
by (auto intro: complete_lattice_class.SUP_upper)  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
747  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
748  | 
lemma suminf_0_le:  | 
| 49664 | 749  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
750  | 
assumes "\<And>n. 0 \<le> f n"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
751  | 
shows "0 \<le> (\<Sum>n. f n)"  | 
| 53788 | 752  | 
using suminf_upper[of f 0, OF assms]  | 
753  | 
by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
754  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
755  | 
lemma suminf_le_pos:  | 
| 43920 | 756  | 
fixes f g :: "nat \<Rightarrow> ereal"  | 
| 53788 | 757  | 
assumes "\<And>N. f N \<le> g N"  | 
758  | 
and "\<And>N. 0 \<le> f N"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
759  | 
shows "suminf f \<le> suminf g"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
760  | 
proof (safe intro!: suminf_bound)  | 
| 49664 | 761  | 
fix n  | 
| 53788 | 762  | 
  {
 | 
763  | 
fix N  | 
|
764  | 
have "0 \<le> g N"  | 
|
765  | 
using assms(2,1)[of N] by auto  | 
|
766  | 
}  | 
|
| 49664 | 767  | 
  have "setsum f {..<n} \<le> setsum g {..<n}"
 | 
768  | 
using assms by (auto intro: setsum_mono)  | 
|
| 53788 | 769  | 
also have "\<dots> \<le> suminf g"  | 
770  | 
using `\<And>N. 0 \<le> g N`  | 
|
771  | 
by (rule suminf_upper)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
772  | 
  finally show "setsum f {..<n} \<le> suminf g" .
 | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
773  | 
qed (rule assms(2))  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
774  | 
|
| 53788 | 775  | 
lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal) ^ Suc n) = 1"  | 
| 43920 | 776  | 
using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric]  | 
777  | 
by (simp add: one_ereal_def)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
778  | 
|
| 43920 | 779  | 
lemma suminf_add_ereal:  | 
780  | 
fixes f g :: "nat \<Rightarrow> ereal"  | 
|
| 53788 | 781  | 
assumes "\<And>i. 0 \<le> f i"  | 
782  | 
and "\<And>i. 0 \<le> g i"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
783  | 
shows "(\<Sum>i. f i + g i) = suminf f + suminf g"  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
784  | 
apply (subst (1 2 3) suminf_ereal_eq_SUP)  | 
| 57418 | 785  | 
unfolding setsum.distrib  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
786  | 
apply (intro assms ereal_add_nonneg_nonneg SUP_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+  | 
| 49664 | 787  | 
done  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
788  | 
|
| 43920 | 789  | 
lemma suminf_cmult_ereal:  | 
790  | 
fixes f g :: "nat \<Rightarrow> ereal"  | 
|
| 53788 | 791  | 
assumes "\<And>i. 0 \<le> f i"  | 
792  | 
and "0 \<le> a"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
793  | 
shows "(\<Sum>i. a * f i) = a * suminf f"  | 
| 43920 | 794  | 
by (auto simp: setsum_ereal_right_distrib[symmetric] assms  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
795  | 
ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUP  | 
| 
59452
 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 
hoelzl 
parents: 
59425 
diff
changeset
 | 
796  | 
intro!: SUP_ereal_mult_left)  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
797  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
798  | 
lemma suminf_PInfty:  | 
| 43923 | 799  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
| 53788 | 800  | 
assumes "\<And>i. 0 \<le> f i"  | 
801  | 
and "suminf f \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
802  | 
shows "f i \<noteq> \<infinity>"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
803  | 
proof -  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
804  | 
from suminf_upper[of f "Suc i", OF assms(1)] assms(2)  | 
| 53788 | 805  | 
have "(\<Sum>i<Suc i. f i) \<noteq> \<infinity>"  | 
806  | 
by auto  | 
|
807  | 
then show ?thesis  | 
|
808  | 
unfolding setsum_Pinfty by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
809  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
810  | 
|
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
811  | 
lemma suminf_PInfty_fun:  | 
| 53788 | 812  | 
assumes "\<And>i. 0 \<le> f i"  | 
813  | 
and "suminf f \<noteq> \<infinity>"  | 
|
| 43920 | 814  | 
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
 | 
815  | 
proof -  | 
| 43920 | 816  | 
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
 | 
817  | 
proof  | 
| 53788 | 818  | 
fix i  | 
819  | 
show "\<exists>r. f i = ereal r"  | 
|
820  | 
using suminf_PInfty[OF assms] assms(1)[of i]  | 
|
821  | 
by (cases "f i") auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
822  | 
qed  | 
| 53788 | 823  | 
from choice[OF this] show ?thesis  | 
824  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
825  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
826  | 
|
| 43920 | 827  | 
lemma summable_ereal:  | 
| 53788 | 828  | 
assumes "\<And>i. 0 \<le> f i"  | 
829  | 
and "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
830  | 
shows "summable f"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
831  | 
proof -  | 
| 43920 | 832  | 
have "0 \<le> (\<Sum>i. ereal (f i))"  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
833  | 
using assms by (intro suminf_0_le) auto  | 
| 43920 | 834  | 
with assms obtain r where r: "(\<Sum>i. ereal (f i)) = ereal r"  | 
835  | 
by (cases "\<Sum>i. ereal (f i)") auto  | 
|
836  | 
from summable_ereal_pos[of "\<lambda>x. ereal (f x)"]  | 
|
| 53788 | 837  | 
have "summable (\<lambda>x. ereal (f x))"  | 
838  | 
using assms by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
839  | 
from summable_sums[OF this]  | 
| 53788 | 840  | 
have "(\<lambda>x. ereal (f x)) sums (\<Sum>x. ereal (f x))"  | 
841  | 
by auto  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
842  | 
then show "summable f"  | 
| 43920 | 843  | 
unfolding r sums_ereal summable_def ..  | 
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
844  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
845  | 
|
| 43920 | 846  | 
lemma suminf_ereal:  | 
| 53788 | 847  | 
assumes "\<And>i. 0 \<le> f i"  | 
848  | 
and "(\<Sum>i. ereal (f i)) \<noteq> \<infinity>"  | 
|
| 43920 | 849  | 
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
 | 
850  | 
proof (rule sums_unique[symmetric])  | 
| 43920 | 851  | 
from summable_ereal[OF assms]  | 
852  | 
show "(\<lambda>x. ereal (f x)) sums (ereal (suminf f))"  | 
|
| 53788 | 853  | 
unfolding sums_ereal  | 
854  | 
using assms  | 
|
855  | 
by (intro summable_sums summable_ereal)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
856  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
857  | 
|
| 43920 | 858  | 
lemma suminf_ereal_minus:  | 
859  | 
fixes f g :: "nat \<Rightarrow> ereal"  | 
|
| 53788 | 860  | 
assumes ord: "\<And>i. g i \<le> f i" "\<And>i. 0 \<le> g i"  | 
861  | 
and fin: "suminf f \<noteq> \<infinity>" "suminf g \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
862  | 
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
 | 
863  | 
proof -  | 
| 53788 | 864  | 
  {
 | 
865  | 
fix i  | 
|
866  | 
have "0 \<le> f i"  | 
|
867  | 
using ord[of i] by auto  | 
|
868  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
869  | 
moreover  | 
| 53788 | 870  | 
from suminf_PInfty_fun[OF `\<And>i. 0 \<le> f i` fin(1)] obtain f' where [simp]: "f = (\<lambda>x. ereal (f' x))" ..  | 
871  | 
from suminf_PInfty_fun[OF `\<And>i. 0 \<le> g i` fin(2)] obtain g' where [simp]: "g = (\<lambda>x. ereal (g' x))" ..  | 
|
872  | 
  {
 | 
|
873  | 
fix i  | 
|
874  | 
have "0 \<le> f i - g i"  | 
|
875  | 
using ord[of i] by (auto simp: ereal_le_minus_iff)  | 
|
876  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
877  | 
moreover  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
878  | 
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
 | 
879  | 
using assms by (auto intro!: suminf_le_pos simp: field_simps)  | 
| 53788 | 880  | 
then have "suminf (\<lambda>i. f i - g i) \<noteq> \<infinity>"  | 
881  | 
using fin by auto  | 
|
882  | 
ultimately show ?thesis  | 
|
883  | 
using assms `\<And>i. 0 \<le> f i`  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
884  | 
apply simp  | 
| 49664 | 885  | 
apply (subst (1 2 3) suminf_ereal)  | 
886  | 
apply (auto intro!: suminf_diff[symmetric] summable_ereal)  | 
|
887  | 
done  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
888  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
889  | 
|
| 49664 | 890  | 
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
 | 
891  | 
proof -  | 
| 53788 | 892  | 
have "(\<Sum>i<Suc 0. \<infinity>) \<le> (\<Sum>x. \<infinity>::ereal)"  | 
893  | 
by (rule suminf_upper) auto  | 
|
894  | 
then show ?thesis  | 
|
895  | 
by simp  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
896  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
897  | 
|
| 43920 | 898  | 
lemma summable_real_of_ereal:  | 
| 43923 | 899  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
| 49664 | 900  | 
assumes f: "\<And>i. 0 \<le> f i"  | 
901  | 
and fin: "(\<Sum>i. f i) \<noteq> \<infinity>"  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
902  | 
shows "summable (\<lambda>i. real (f i))"  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
903  | 
proof (rule summable_def[THEN iffD2])  | 
| 53788 | 904  | 
have "0 \<le> (\<Sum>i. f i)"  | 
905  | 
using assms by (auto intro: suminf_0_le)  | 
|
906  | 
with fin obtain r where r: "ereal r = (\<Sum>i. f i)"  | 
|
907  | 
by (cases "(\<Sum>i. f i)") auto  | 
|
908  | 
  {
 | 
|
909  | 
fix i  | 
|
910  | 
have "f i \<noteq> \<infinity>"  | 
|
911  | 
using f by (intro suminf_PInfty[OF _ fin]) auto  | 
|
912  | 
then have "\<bar>f i\<bar> \<noteq> \<infinity>"  | 
|
913  | 
using f[of i] by auto  | 
|
914  | 
}  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
915  | 
note fin = this  | 
| 43920 | 916  | 
have "(\<lambda>i. ereal (real (f i))) sums (\<Sum>i. ereal (real (f i)))"  | 
| 53788 | 917  | 
using f  | 
| 57865 | 918  | 
by (auto intro!: summable_ereal_pos simp: ereal_le_real_iff zero_ereal_def)  | 
| 53788 | 919  | 
also have "\<dots> = ereal r"  | 
920  | 
using fin r by (auto simp: ereal_real)  | 
|
921  | 
finally show "\<exists>r. (\<lambda>i. real (f i)) sums r"  | 
|
922  | 
by (auto simp: sums_ereal)  | 
|
| 
41980
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
923  | 
qed  | 
| 
 
28b51effc5ed
split Extended_Reals into parts for Library and Multivariate_Analysis
 
hoelzl 
parents:  
diff
changeset
 | 
924  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
925  | 
lemma suminf_SUP_eq:  | 
| 43920 | 926  | 
fixes f :: "nat \<Rightarrow> nat \<Rightarrow> ereal"  | 
| 53788 | 927  | 
assumes "\<And>i. incseq (\<lambda>n. f n i)"  | 
928  | 
and "\<And>n i. 0 \<le> f n i"  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
929  | 
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: 
41983 
diff
changeset
 | 
930  | 
proof -  | 
| 53788 | 931  | 
  {
 | 
932  | 
fix n :: nat  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
933  | 
have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"  | 
| 53788 | 934  | 
using assms  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
935  | 
by (auto intro!: SUP_ereal_setsum [symmetric])  | 
| 53788 | 936  | 
}  | 
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
937  | 
note * = this  | 
| 53788 | 938  | 
show ?thesis  | 
939  | 
using assms  | 
|
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
940  | 
apply (subst (1 2) suminf_ereal_eq_SUP)  | 
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
941  | 
unfolding *  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
942  | 
apply (auto intro!: SUP_upper2)  | 
| 49664 | 943  | 
apply (subst SUP_commute)  | 
944  | 
apply rule  | 
|
945  | 
done  | 
|
| 
42950
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
946  | 
qed  | 
| 
 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 
hoelzl 
parents: 
41983 
diff
changeset
 | 
947  | 
|
| 47761 | 948  | 
lemma suminf_setsum_ereal:  | 
949  | 
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> ereal"  | 
|
950  | 
assumes nonneg: "\<And>i a. a \<in> A \<Longrightarrow> 0 \<le> f i a"  | 
|
951  | 
shows "(\<Sum>i. \<Sum>a\<in>A. f i a) = (\<Sum>a\<in>A. \<Sum>i. f i a)"  | 
|
| 53788 | 952  | 
proof (cases "finite A")  | 
953  | 
case True  | 
|
954  | 
then show ?thesis  | 
|
955  | 
using nonneg  | 
|
| 47761 | 956  | 
by induct (simp_all add: suminf_add_ereal setsum_nonneg)  | 
| 53788 | 957  | 
next  | 
958  | 
case False  | 
|
959  | 
then show ?thesis by simp  | 
|
960  | 
qed  | 
|
| 47761 | 961  | 
|
| 50104 | 962  | 
lemma suminf_ereal_eq_0:  | 
963  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
|
964  | 
assumes nneg: "\<And>i. 0 \<le> f i"  | 
|
965  | 
shows "(\<Sum>i. f i) = 0 \<longleftrightarrow> (\<forall>i. f i = 0)"  | 
|
966  | 
proof  | 
|
967  | 
assume "(\<Sum>i. f i) = 0"  | 
|
| 53788 | 968  | 
  {
 | 
969  | 
fix i  | 
|
970  | 
assume "f i \<noteq> 0"  | 
|
971  | 
with nneg have "0 < f i"  | 
|
972  | 
by (auto simp: less_le)  | 
|
| 50104 | 973  | 
also have "f i = (\<Sum>j. if j = i then f i else 0)"  | 
974  | 
      by (subst suminf_finite[where N="{i}"]) auto
 | 
|
975  | 
also have "\<dots> \<le> (\<Sum>i. f i)"  | 
|
| 53788 | 976  | 
using nneg  | 
977  | 
by (auto intro!: suminf_le_pos)  | 
|
978  | 
finally have False  | 
|
979  | 
using `(\<Sum>i. f i) = 0` by auto  | 
|
980  | 
}  | 
|
981  | 
then show "\<forall>i. f i = 0"  | 
|
982  | 
by auto  | 
|
| 50104 | 983  | 
qed simp  | 
984  | 
||
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
985  | 
lemma Liminf_within:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
986  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
987  | 
  shows "Liminf (at x within S) f = (SUP e:{0<..}. INF y:(S \<inter> ball x e - {x}). f y)"
 | 
| 
51641
 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 
hoelzl 
parents: 
51530 
diff
changeset
 | 
988  | 
unfolding Liminf_def eventually_at  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
989  | 
proof (rule SUP_eq, simp_all add: Ball_def Bex_def, safe)  | 
| 53788 | 990  | 
fix P d  | 
991  | 
assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
992  | 
  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: 
51329 
diff
changeset
 | 
993  | 
by (auto simp: zero_less_dist_iff dist_commute)  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
994  | 
  then show "\<exists>r>0. INFIMUM (Collect P) f \<le> INFIMUM (S \<inter> ball x r - {x}) f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
995  | 
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: 
51329 
diff
changeset
 | 
996  | 
next  | 
| 53788 | 997  | 
fix d :: real  | 
998  | 
assume "0 < d"  | 
|
| 
51641
 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 
hoelzl 
parents: 
51530 
diff
changeset
 | 
999  | 
then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> xa \<noteq> x \<and> dist xa x < d \<longrightarrow> P xa) \<and>  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1000  | 
    INFIMUM (S \<inter> ball x d - {x}) f \<le> INFIMUM (Collect P) f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1001  | 
    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: 
51329 
diff
changeset
 | 
1002  | 
(auto intro!: INF_mono exI[of _ d] simp: dist_commute)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1003  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1004  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1005  | 
lemma Limsup_within:  | 
| 53788 | 1006  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1007  | 
  shows "Limsup (at x within S) f = (INF e:{0<..}. SUP y:(S \<inter> ball x e - {x}). f y)"
 | 
| 
51641
 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 
hoelzl 
parents: 
51530 
diff
changeset
 | 
1008  | 
unfolding Limsup_def eventually_at  | 
| 
56212
 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 
haftmann 
parents: 
56166 
diff
changeset
 | 
1009  | 
proof (rule INF_eq, simp_all add: Ball_def Bex_def, safe)  | 
| 53788 | 1010  | 
fix P d  | 
1011  | 
assume "0 < d" and "\<forall>y. y \<in> S \<longrightarrow> y \<noteq> x \<and> dist y x < d \<longrightarrow> P y"  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1012  | 
  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: 
51329 
diff
changeset
 | 
1013  | 
by (auto simp: zero_less_dist_iff dist_commute)  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1014  | 
  then show "\<exists>r>0. SUPREMUM (S \<inter> ball x r - {x}) f \<le> SUPREMUM (Collect P) f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1015  | 
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: 
51329 
diff
changeset
 | 
1016  | 
next  | 
| 53788 | 1017  | 
fix d :: real  | 
1018  | 
assume "0 < d"  | 
|
| 
51641
 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 
hoelzl 
parents: 
51530 
diff
changeset
 | 
1019  | 
then show "\<exists>P. (\<exists>d>0. \<forall>xa. xa \<in> S \<longrightarrow> xa \<noteq> x \<and> dist xa x < d \<longrightarrow> P xa) \<and>  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1020  | 
    SUPREMUM (Collect P) f \<le> SUPREMUM (S \<inter> ball x d - {x}) f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1021  | 
    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: 
51329 
diff
changeset
 | 
1022  | 
(auto intro!: SUP_mono exI[of _ d] simp: dist_commute)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1023  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1024  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1025  | 
lemma Liminf_at:  | 
| 
54257
 
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
 
hoelzl 
parents: 
53788 
diff
changeset
 | 
1026  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1027  | 
  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: 
51329 
diff
changeset
 | 
1028  | 
using Liminf_within[of x UNIV f] by simp  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1029  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1030  | 
lemma Limsup_at:  | 
| 
54257
 
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
 
hoelzl 
parents: 
53788 
diff
changeset
 | 
1031  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1032  | 
  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: 
51329 
diff
changeset
 | 
1033  | 
using Limsup_within[of x UNIV f] by simp  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1034  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1035  | 
lemma min_Liminf_at:  | 
| 53788 | 1036  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_linorder"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1037  | 
  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: 
51329 
diff
changeset
 | 
1038  | 
unfolding inf_min[symmetric] Liminf_at  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1039  | 
apply (subst inf_commute)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1040  | 
apply (subst SUP_inf)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1041  | 
apply (intro SUP_cong[OF refl])  | 
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54258 
diff
changeset
 | 
1042  | 
  apply (cut_tac A="ball x xa - {x}" and B="{x}" and M=f in INF_union)
 | 
| 56166 | 1043  | 
apply (drule sym)  | 
1044  | 
apply auto  | 
|
| 57865 | 1045  | 
apply (metis INF_absorb centre_in_ball)  | 
1046  | 
done  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1047  | 
|
| 53788 | 1048  | 
|
| 57025 | 1049  | 
lemma suminf_ereal_offset_le:  | 
1050  | 
fixes f :: "nat \<Rightarrow> ereal"  | 
|
1051  | 
assumes f: "\<And>i. 0 \<le> f i"  | 
|
1052  | 
shows "(\<Sum>i. f (i + k)) \<le> suminf f"  | 
|
1053  | 
proof -  | 
|
1054  | 
have "(\<lambda>n. \<Sum>i<n. f (i + k)) ----> (\<Sum>i. f (i + k))"  | 
|
1055  | 
using summable_sums[OF summable_ereal_pos] by (simp add: sums_def atLeast0LessThan f)  | 
|
1056  | 
moreover have "(\<lambda>n. \<Sum>i<n. f i) ----> (\<Sum>i. f i)"  | 
|
1057  | 
using summable_sums[OF summable_ereal_pos] by (simp add: sums_def atLeast0LessThan f)  | 
|
1058  | 
then have "(\<lambda>n. \<Sum>i<n + k. f i) ----> (\<Sum>i. f i)"  | 
|
1059  | 
by (rule LIMSEQ_ignore_initial_segment)  | 
|
1060  | 
ultimately show ?thesis  | 
|
1061  | 
proof (rule LIMSEQ_le, safe intro!: exI[of _ k])  | 
|
1062  | 
fix n assume "k \<le> n"  | 
|
1063  | 
have "(\<Sum>i<n. f (i + k)) = (\<Sum>i<n. (f \<circ> (\<lambda>i. i + k)) i)"  | 
|
1064  | 
by simp  | 
|
1065  | 
    also have "\<dots> = (\<Sum>i\<in>(\<lambda>i. i + k) ` {..<n}. f i)"
 | 
|
| 57418 | 1066  | 
by (subst setsum.reindex) auto  | 
| 57025 | 1067  | 
    also have "\<dots> \<le> setsum f {..<n + k}"
 | 
1068  | 
by (intro setsum_mono3) (auto simp: f)  | 
|
1069  | 
    finally show "(\<Sum>i<n. f (i + k)) \<le> setsum f {..<n + k}" .
 | 
|
1070  | 
qed  | 
|
1071  | 
qed  | 
|
1072  | 
||
1073  | 
lemma sums_suminf_ereal: "f sums x \<Longrightarrow> (\<Sum>i. ereal (f i)) = ereal x"  | 
|
1074  | 
by (metis sums_ereal sums_unique)  | 
|
1075  | 
||
1076  | 
lemma suminf_ereal': "summable f \<Longrightarrow> (\<Sum>i. ereal (f i)) = ereal (\<Sum>i. f i)"  | 
|
1077  | 
by (metis sums_ereal sums_unique summable_def)  | 
|
1078  | 
||
1079  | 
lemma suminf_ereal_finite: "summable f \<Longrightarrow> (\<Sum>i. ereal (f i)) \<noteq> \<infinity>"  | 
|
1080  | 
by (auto simp: sums_ereal[symmetric] summable_def sums_unique[symmetric])  | 
|
1081  | 
||
| 59425 | 1082  | 
lemma suminf_ereal_finite_neg:  | 
1083  | 
assumes "summable f"  | 
|
1084  | 
shows "(\<Sum>x. ereal (f x)) \<noteq> -\<infinity>"  | 
|
1085  | 
proof-  | 
|
1086  | 
from assms obtain x where "f sums x" by blast  | 
|
1087  | 
hence "(\<lambda>x. ereal (f x)) sums ereal x" by (simp add: sums_ereal)  | 
|
1088  | 
from sums_unique[OF this] have "(\<Sum>x. ereal (f x)) = ereal x" ..  | 
|
1089  | 
thus "(\<Sum>x. ereal (f x)) \<noteq> -\<infinity>" by simp_all  | 
|
1090  | 
qed  | 
|
1091  | 
||
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1092  | 
subsection {* monoset *}
 | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1093  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1094  | 
definition (in order) mono_set:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1095  | 
"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: 
51329 
diff
changeset
 | 
1096  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1097  | 
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: 
51329 
diff
changeset
 | 
1098  | 
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: 
51329 
diff
changeset
 | 
1099  | 
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: 
51329 
diff
changeset
 | 
1100  | 
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: 
51329 
diff
changeset
 | 
1101  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1102  | 
lemma (in complete_linorder) mono_set_iff:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1103  | 
fixes S :: "'a set"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1104  | 
defines "a \<equiv> Inf S"  | 
| 53788 | 1105  | 
  shows "mono_set S \<longleftrightarrow> S = {a <..} \<or> S = {a..}" (is "_ = ?c")
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1106  | 
proof  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1107  | 
assume "mono_set S"  | 
| 53788 | 1108  | 
then have mono: "\<And>x y. x \<le> y \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S"  | 
1109  | 
by (auto simp: mono_set)  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1110  | 
show ?c  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1111  | 
proof cases  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1112  | 
assume "a \<in> S"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1113  | 
show ?c  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1114  | 
using mono[OF _ `a \<in> S`]  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1115  | 
by (auto intro: Inf_lower simp: a_def)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1116  | 
next  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1117  | 
assume "a \<notin> S"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1118  | 
    have "S = {a <..}"
 | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1119  | 
proof safe  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1120  | 
fix x assume "x \<in> S"  | 
| 53788 | 1121  | 
then have "a \<le> x"  | 
1122  | 
unfolding a_def by (rule Inf_lower)  | 
|
1123  | 
then show "a < x"  | 
|
1124  | 
using `x \<in> S` `a \<notin> S` by (cases "a = x") auto  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1125  | 
next  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1126  | 
fix x assume "a < x"  | 
| 53788 | 1127  | 
then obtain y where "y < x" "y \<in> S"  | 
1128  | 
unfolding a_def Inf_less_iff ..  | 
|
1129  | 
with mono[of y x] show "x \<in> S"  | 
|
1130  | 
by auto  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1131  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1132  | 
then show ?c ..  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1133  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1134  | 
qed auto  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1135  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1136  | 
lemma ereal_open_mono_set:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1137  | 
fixes S :: "ereal set"  | 
| 53788 | 1138  | 
  shows "open S \<and> mono_set S \<longleftrightarrow> S = UNIV \<or> S = {Inf S <..}"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1139  | 
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: 
51329 
diff
changeset
 | 
1140  | 
ereal_open_closed mono_set_iff open_ereal_greaterThan)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1141  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1142  | 
lemma ereal_closed_mono_set:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1143  | 
fixes S :: "ereal set"  | 
| 53788 | 1144  | 
  shows "closed S \<and> mono_set S \<longleftrightarrow> S = {} \<or> S = {Inf S ..}"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1145  | 
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: 
51329 
diff
changeset
 | 
1146  | 
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: 
51329 
diff
changeset
 | 
1147  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1148  | 
lemma ereal_Liminf_Sup_monoset:  | 
| 53788 | 1149  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1150  | 
shows "Liminf net f =  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1151  | 
    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: 
51329 
diff
changeset
 | 
1152  | 
(is "_ = Sup ?A")  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1153  | 
proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least)  | 
| 53788 | 1154  | 
fix P  | 
1155  | 
assume P: "eventually P net"  | 
|
1156  | 
fix S  | 
|
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1157  | 
assume S: "mono_set S" "INFIMUM (Collect P) f \<in> S"  | 
| 53788 | 1158  | 
  {
 | 
1159  | 
fix x  | 
|
1160  | 
assume "P x"  | 
|
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1161  | 
then have "INFIMUM (Collect P) f \<le> f x"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1162  | 
by (intro complete_lattice_class.INF_lower) simp  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1163  | 
with S have "f x \<in> S"  | 
| 53788 | 1164  | 
by (simp add: mono_set)  | 
1165  | 
}  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1166  | 
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: 
51329 
diff
changeset
 | 
1167  | 
by (auto elim: eventually_elim1)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1168  | 
next  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1169  | 
fix y l  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1170  | 
assume S: "\<forall>S. open S \<longrightarrow> mono_set S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net"  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1171  | 
assume P: "\<forall>P. eventually P net \<longrightarrow> INFIMUM (Collect P) f \<le> y"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1172  | 
show "l \<le> y"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1173  | 
proof (rule dense_le)  | 
| 53788 | 1174  | 
fix B  | 
1175  | 
assume "B < l"  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1176  | 
    then have "eventually (\<lambda>x. f x \<in> {B <..}) net"
 | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1177  | 
by (intro S[rule_format]) auto  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1178  | 
    then have "INFIMUM {x. B < f x} f \<le> y"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1179  | 
using P by auto  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1180  | 
    moreover have "B \<le> INFIMUM {x. B < f x} f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1181  | 
by (intro INF_greatest) auto  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1182  | 
ultimately show "B \<le> y"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1183  | 
by simp  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1184  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1185  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1186  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1187  | 
lemma ereal_Limsup_Inf_monoset:  | 
| 53788 | 1188  | 
fixes f :: "'a \<Rightarrow> ereal"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1189  | 
shows "Limsup net f =  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1190  | 
    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: 
51329 
diff
changeset
 | 
1191  | 
(is "_ = Inf ?A")  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1192  | 
proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest)  | 
| 53788 | 1193  | 
fix P  | 
1194  | 
assume P: "eventually P net"  | 
|
1195  | 
fix S  | 
|
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1196  | 
assume S: "mono_set (uminus`S)" "SUPREMUM (Collect P) f \<in> S"  | 
| 53788 | 1197  | 
  {
 | 
1198  | 
fix x  | 
|
1199  | 
assume "P x"  | 
|
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1200  | 
then have "f x \<le> SUPREMUM (Collect P) f"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1201  | 
by (intro complete_lattice_class.SUP_upper) simp  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1202  | 
with S(1)[unfolded mono_set, rule_format, of "- SUPREMUM (Collect P) f" "- f x"] S(2)  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1203  | 
have "f x \<in> S"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1204  | 
by (simp add: inj_image_mem_iff) }  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1205  | 
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: 
51329 
diff
changeset
 | 
1206  | 
by (auto elim: eventually_elim1)  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1207  | 
next  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1208  | 
fix y l  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1209  | 
assume S: "\<forall>S. open S \<longrightarrow> mono_set (uminus ` S) \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net"  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1210  | 
assume P: "\<forall>P. eventually P net \<longrightarrow> y \<le> SUPREMUM (Collect P) f"  | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1211  | 
show "y \<le> l"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1212  | 
proof (rule dense_ge)  | 
| 53788 | 1213  | 
fix B  | 
1214  | 
assume "l < B"  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1215  | 
    then have "eventually (\<lambda>x. f x \<in> {..< B}) net"
 | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1216  | 
by (intro S[rule_format]) auto  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1217  | 
    then have "y \<le> SUPREMUM {x. f x < B} f"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1218  | 
using P by auto  | 
| 
56218
 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 
haftmann 
parents: 
56212 
diff
changeset
 | 
1219  | 
    moreover have "SUPREMUM {x. f x < B} f \<le> B"
 | 
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1220  | 
by (intro SUP_least) auto  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1221  | 
ultimately show "y \<le> B"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1222  | 
by simp  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1223  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1224  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1225  | 
|
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1226  | 
lemma liminf_bounded_open:  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1227  | 
fixes x :: "nat \<Rightarrow> ereal"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1228  | 
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: 
51329 
diff
changeset
 | 
1229  | 
(is "_ \<longleftrightarrow> ?P x0")  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1230  | 
proof  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1231  | 
assume "?P x0"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1232  | 
then show "x0 \<le> liminf x"  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1233  | 
unfolding ereal_Liminf_Sup_monoset eventually_sequentially  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1234  | 
by (intro complete_lattice_class.Sup_upper) auto  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1235  | 
next  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1236  | 
assume "x0 \<le> liminf x"  | 
| 53788 | 1237  | 
  {
 | 
1238  | 
fix S :: "ereal set"  | 
|
1239  | 
assume om: "open S" "mono_set S" "x0 \<in> S"  | 
|
1240  | 
    {
 | 
|
1241  | 
assume "S = UNIV"  | 
|
1242  | 
then have "\<exists>N. \<forall>n\<ge>N. x n \<in> S"  | 
|
1243  | 
by auto  | 
|
1244  | 
}  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1245  | 
moreover  | 
| 53788 | 1246  | 
    {
 | 
1247  | 
assume "S \<noteq> UNIV"  | 
|
1248  | 
      then obtain B where B: "S = {B<..}"
 | 
|
1249  | 
using om ereal_open_mono_set by auto  | 
|
1250  | 
then have "B < x0"  | 
|
1251  | 
using om by auto  | 
|
1252  | 
then have "\<exists>N. \<forall>n\<ge>N. x n \<in> S"  | 
|
1253  | 
unfolding B  | 
|
1254  | 
using `x0 \<le> liminf x` liminf_bounded_iff  | 
|
1255  | 
by auto  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1256  | 
}  | 
| 53788 | 1257  | 
ultimately have "\<exists>N. \<forall>n\<ge>N. x n \<in> S"  | 
1258  | 
by auto  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1259  | 
}  | 
| 53788 | 1260  | 
then show "?P x0"  | 
1261  | 
by auto  | 
|
| 
51340
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1262  | 
qed  | 
| 
 
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
 
hoelzl 
parents: 
51329 
diff
changeset
 | 
1263  | 
|
| 
57446
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1264  | 
subsection "Relate extended reals and the indicator function"  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1265  | 
|
| 59000 | 1266  | 
lemma ereal_indicator_le_0: "(indicator S x::ereal) \<le> 0 \<longleftrightarrow> x \<notin> S"  | 
1267  | 
by (auto split: split_indicator simp: one_ereal_def)  | 
|
1268  | 
||
| 
57446
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1269  | 
lemma ereal_indicator: "ereal (indicator A x) = indicator A x"  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1270  | 
by (auto simp: indicator_def one_ereal_def)  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1271  | 
|
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1272  | 
lemma ereal_mult_indicator: "ereal (x * indicator A y) = ereal x * indicator A y"  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1273  | 
by (simp split: split_indicator)  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1274  | 
|
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1275  | 
lemma ereal_indicator_mult: "ereal (indicator A y * x) = indicator A y * ereal x"  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1276  | 
by (simp split: split_indicator)  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1277  | 
|
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1278  | 
lemma ereal_indicator_nonneg[simp, intro]: "0 \<le> (indicator A x ::ereal)"  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1279  | 
unfolding indicator_def by auto  | 
| 
 
06e195515deb
some lemmas about the indicator function; removed lemma sums_def2
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
1280  | 
|
| 59425 | 1281  | 
lemma indicator_inter_arith_ereal: "indicator A x * indicator B x = (indicator (A \<inter> B) x :: ereal)"  | 
1282  | 
by (simp split: split_indicator)  | 
|
1283  | 
||
1284  | 
||
| 44125 | 1285  | 
end  |