author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 70378 | ebd108578ab1 |
child 81332 | f94b30fa2b6c |
permissions | -rw-r--r-- |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
1 |
(* Title: HOL/Library/Liminf_Limsup.thy |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
2 |
Author: Johannes Hölzl, TU München |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
3 |
Author: Manuel Eberl, TU München |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
4 |
*) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
5 |
|
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
6 |
section \<open>Liminf and Limsup on conditionally complete lattices\<close> |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
7 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
8 |
theory Liminf_Limsup |
51542 | 9 |
imports Complex_Main |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
10 |
begin |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
11 |
|
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
12 |
lemma (in conditionally_complete_linorder) le_cSup_iff: |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
13 |
assumes "A \<noteq> {}" "bdd_above A" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
14 |
shows "x \<le> Sup A \<longleftrightarrow> (\<forall>y<x. \<exists>a\<in>A. y < a)" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
15 |
proof safe |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
16 |
fix y assume "x \<le> Sup A" "y < x" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
17 |
then have "y < Sup A" by auto |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
18 |
then show "\<exists>a\<in>A. y < a" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
19 |
unfolding less_cSup_iff[OF assms] . |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
20 |
qed (auto elim!: allE[of _ "Sup A"] simp add: not_le[symmetric] cSup_upper assms) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
21 |
|
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
22 |
lemma (in conditionally_complete_linorder) le_cSUP_iff: |
69313 | 23 |
"A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> x \<le> Sup (f ` A) \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y < f i)" |
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
24 |
using le_cSup_iff [of "f ` A"] by simp |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
25 |
|
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
26 |
lemma le_cSup_iff_less: |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
27 |
fixes x :: "'a :: {conditionally_complete_linorder, dense_linorder}" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
28 |
shows "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> x \<le> (SUP i\<in>A. f i) \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y \<le> f i)" |
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
29 |
by (simp add: le_cSUP_iff) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
30 |
(blast intro: less_imp_le less_trans less_le_trans dest: dense) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
31 |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
32 |
lemma le_Sup_iff_less: |
53216 | 33 |
fixes x :: "'a :: {complete_linorder, dense_linorder}" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
34 |
shows "x \<le> (SUP i\<in>A. f i) \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y \<le> f i)" (is "?lhs = ?rhs") |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
35 |
unfolding le_SUP_iff |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
36 |
by (blast intro: less_imp_le less_trans less_le_trans dest: dense) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
37 |
|
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
38 |
lemma (in conditionally_complete_linorder) cInf_le_iff: |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
39 |
assumes "A \<noteq> {}" "bdd_below A" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
40 |
shows "Inf A \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>a\<in>A. y > a)" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
41 |
proof safe |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
42 |
fix y assume "x \<ge> Inf A" "y > x" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
43 |
then have "y > Inf A" by auto |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
44 |
then show "\<exists>a\<in>A. y > a" |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
45 |
unfolding cInf_less_iff[OF assms] . |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
46 |
qed (auto elim!: allE[of _ "Inf A"] simp add: not_le[symmetric] cInf_lower assms) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
47 |
|
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
48 |
lemma (in conditionally_complete_linorder) cINF_le_iff: |
69313 | 49 |
"A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> Inf (f ` A) \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. y > f i)" |
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
50 |
using cInf_le_iff [of "f ` A"] by simp |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
51 |
|
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
52 |
lemma cInf_le_iff_less: |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
53 |
fixes x :: "'a :: {conditionally_complete_linorder, dense_linorder}" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
54 |
shows "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> (INF i\<in>A. f i) \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. f i \<le> y)" |
62624
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
55 |
by (simp add: cINF_le_iff) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
56 |
(blast intro: less_imp_le less_trans le_less_trans dest: dense) |
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
hoelzl
parents:
62343
diff
changeset
|
57 |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
58 |
lemma Inf_le_iff_less: |
53216 | 59 |
fixes x :: "'a :: {complete_linorder, dense_linorder}" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
60 |
shows "(INF i\<in>A. f i) \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. f i \<le> y)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
61 |
unfolding INF_le_iff |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
62 |
by (blast intro: less_imp_le less_trans le_less_trans dest: dense) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
63 |
|
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54261
diff
changeset
|
64 |
lemma SUP_pair: |
54257
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents:
53381
diff
changeset
|
65 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
66 |
shows "(SUP i \<in> A. SUP j \<in> B. f i j) = (SUP p \<in> A \<times> B. f (fst p) (snd p))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
67 |
by (rule antisym) (auto intro!: SUP_least SUP_upper2) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
68 |
|
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54261
diff
changeset
|
69 |
lemma INF_pair: |
54257
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents:
53381
diff
changeset
|
70 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
71 |
shows "(INF i \<in> A. INF j \<in> B. f i j) = (INF p \<in> A \<times> B. f (fst p) (snd p))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
72 |
by (rule antisym) (auto intro!: INF_greatest INF_lower2) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
73 |
|
68860
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
74 |
lemma INF_Sigma: |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
75 |
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
76 |
shows "(INF i \<in> A. INF j \<in> B i. f i j) = (INF p \<in> Sigma A B. f (fst p) (snd p))" |
68860
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
77 |
by (rule antisym) (auto intro!: INF_greatest INF_lower2) |
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
78 |
|
61585 | 79 |
subsubsection \<open>\<open>Liminf\<close> and \<open>Limsup\<close>\<close> |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
80 |
|
54261 | 81 |
definition Liminf :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
82 |
"Liminf F f = (SUP P\<in>{P. eventually P F}. INF x\<in>{x. P x}. f x)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
83 |
|
54261 | 84 |
definition Limsup :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
85 |
"Limsup F f = (INF P\<in>{P. eventually P F}. SUP x\<in>{x. P x}. f x)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
86 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
87 |
abbreviation "liminf \<equiv> Liminf sequentially" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
88 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
89 |
abbreviation "limsup \<equiv> Limsup sequentially" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
90 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
91 |
lemma Liminf_eqI: |
69313 | 92 |
"(\<And>P. eventually P F \<Longrightarrow> Inf (f ` (Collect P)) \<le> x) \<Longrightarrow> |
93 |
(\<And>y. (\<And>P. eventually P F \<Longrightarrow> Inf (f ` (Collect P)) \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> Liminf F f = x" |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
94 |
unfolding Liminf_def by (auto intro!: SUP_eqI) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
95 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
96 |
lemma Limsup_eqI: |
69313 | 97 |
"(\<And>P. eventually P F \<Longrightarrow> x \<le> Sup (f ` (Collect P))) \<Longrightarrow> |
98 |
(\<And>y. (\<And>P. eventually P F \<Longrightarrow> y \<le> Sup (f ` (Collect P))) \<Longrightarrow> y \<le> x) \<Longrightarrow> Limsup F f = x" |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
99 |
unfolding Limsup_def by (auto intro!: INF_eqI) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
100 |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
101 |
lemma liminf_SUP_INF: "liminf f = (SUP n. INF m\<in>{n..}. f m)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
102 |
unfolding Liminf_def eventually_sequentially |
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54261
diff
changeset
|
103 |
by (rule SUP_eq) (auto simp: atLeast_def intro!: INF_mono) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
104 |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
105 |
lemma limsup_INF_SUP: "limsup f = (INF n. SUP m\<in>{n..}. f m)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
106 |
unfolding Limsup_def eventually_sequentially |
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54261
diff
changeset
|
107 |
by (rule INF_eq) (auto simp: atLeast_def intro!: SUP_mono) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
108 |
|
70378
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
109 |
lemma mem_limsup_iff: "x \<in> limsup A \<longleftrightarrow> (\<exists>\<^sub>F n in sequentially. x \<in> A n)" |
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
110 |
by (simp add: Limsup_def) (metis (mono_tags) eventually_mono not_frequently) |
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
111 |
|
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
112 |
lemma mem_liminf_iff: "x \<in> liminf A \<longleftrightarrow> (\<forall>\<^sub>F n in sequentially. x \<in> A n)" |
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
113 |
by (simp add: Liminf_def) (metis (mono_tags) eventually_mono) |
ebd108578ab1
more new material about analysis
paulson <lp15@cam.ac.uk>
parents:
69861
diff
changeset
|
114 |
|
61730 | 115 |
lemma Limsup_const: |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
116 |
assumes ntriv: "\<not> trivial_limit F" |
54261 | 117 |
shows "Limsup F (\<lambda>x. c) = c" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
118 |
proof - |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
119 |
have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
120 |
have "\<And>P. eventually P F \<Longrightarrow> (SUP x \<in> {x. P x}. c) = c" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
121 |
using ntriv by (intro SUP_const) (auto simp: eventually_False *) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
122 |
then show ?thesis |
69661 | 123 |
apply (auto simp add: Limsup_def) |
124 |
apply (rule INF_const) |
|
125 |
apply auto |
|
126 |
using eventually_True apply blast |
|
127 |
done |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
128 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
129 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
130 |
lemma Liminf_const: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
131 |
assumes ntriv: "\<not> trivial_limit F" |
54261 | 132 |
shows "Liminf F (\<lambda>x. c) = c" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
133 |
proof - |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
134 |
have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
135 |
have "\<And>P. eventually P F \<Longrightarrow> (INF x \<in> {x. P x}. c) = c" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
136 |
using ntriv by (intro INF_const) (auto simp: eventually_False *) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
137 |
then show ?thesis |
69661 | 138 |
apply (auto simp add: Liminf_def) |
139 |
apply (rule SUP_const) |
|
140 |
apply auto |
|
141 |
using eventually_True apply blast |
|
142 |
done |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
143 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
144 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
145 |
lemma Liminf_mono: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
146 |
assumes ev: "eventually (\<lambda>x. f x \<le> g x) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
147 |
shows "Liminf F f \<le> Liminf F g" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
148 |
unfolding Liminf_def |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
149 |
proof (safe intro!: SUP_mono) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
150 |
fix P assume "eventually P F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
151 |
with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj) |
69313 | 152 |
then show "\<exists>Q\<in>{P. eventually P F}. Inf (f ` (Collect P)) \<le> Inf (g ` (Collect Q))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
153 |
by (intro bexI[of _ ?Q]) (auto intro!: INF_mono) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
154 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
155 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
156 |
lemma Liminf_eq: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
157 |
assumes "eventually (\<lambda>x. f x = g x) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
158 |
shows "Liminf F f = Liminf F g" |
61810 | 159 |
by (intro antisym Liminf_mono eventually_mono[OF assms]) auto |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
160 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
161 |
lemma Limsup_mono: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
162 |
assumes ev: "eventually (\<lambda>x. f x \<le> g x) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
163 |
shows "Limsup F f \<le> Limsup F g" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
164 |
unfolding Limsup_def |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
165 |
proof (safe intro!: INF_mono) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
166 |
fix P assume "eventually P F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
167 |
with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj) |
69313 | 168 |
then show "\<exists>Q\<in>{P. eventually P F}. Sup (f ` (Collect Q)) \<le> Sup (g ` (Collect P))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
169 |
by (intro bexI[of _ ?Q]) (auto intro!: SUP_mono) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
170 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
171 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
172 |
lemma Limsup_eq: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
173 |
assumes "eventually (\<lambda>x. f x = g x) net" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
174 |
shows "Limsup net f = Limsup net g" |
61810 | 175 |
by (intro antisym Limsup_mono eventually_mono[OF assms]) auto |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
176 |
|
63895 | 177 |
lemma Liminf_bot[simp]: "Liminf bot f = top" |
178 |
unfolding Liminf_def top_unique[symmetric] |
|
179 |
by (rule SUP_upper2[where i="\<lambda>x. False"]) simp_all |
|
180 |
||
181 |
lemma Limsup_bot[simp]: "Limsup bot f = bot" |
|
182 |
unfolding Limsup_def bot_unique[symmetric] |
|
183 |
by (rule INF_lower2[where i="\<lambda>x. False"]) simp_all |
|
184 |
||
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
185 |
lemma Liminf_le_Limsup: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
186 |
assumes ntriv: "\<not> trivial_limit F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
187 |
shows "Liminf F f \<le> Limsup F f" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
188 |
unfolding Limsup_def Liminf_def |
54261 | 189 |
apply (rule SUP_least) |
190 |
apply (rule INF_greatest) |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
191 |
proof safe |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
192 |
fix P Q assume "eventually P F" "eventually Q F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
193 |
then have "eventually (\<lambda>x. P x \<and> Q x) F" (is "eventually ?C F") by (rule eventually_conj) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
194 |
then have not_False: "(\<lambda>x. P x \<and> Q x) \<noteq> (\<lambda>x. False)" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
195 |
using ntriv by (auto simp add: eventually_False) |
69313 | 196 |
have "Inf (f ` (Collect P)) \<le> Inf (f ` (Collect ?C))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
197 |
by (rule INF_mono) auto |
69313 | 198 |
also have "\<dots> \<le> Sup (f ` (Collect ?C))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
199 |
using not_False by (intro INF_le_SUP) auto |
69313 | 200 |
also have "\<dots> \<le> Sup (f ` (Collect Q))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
201 |
by (rule SUP_mono) auto |
69313 | 202 |
finally show "Inf (f ` (Collect P)) \<le> Sup (f ` (Collect Q))" . |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
203 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
204 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
205 |
lemma Liminf_bounded: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
206 |
assumes le: "eventually (\<lambda>n. C \<le> X n) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
207 |
shows "C \<le> Liminf F X" |
63895 | 208 |
using Liminf_mono[OF le] Liminf_const[of F C] |
209 |
by (cases "F = bot") simp_all |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
210 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
211 |
lemma Limsup_bounded: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
212 |
assumes le: "eventually (\<lambda>n. X n \<le> C) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
213 |
shows "Limsup F X \<le> C" |
63895 | 214 |
using Limsup_mono[OF le] Limsup_const[of F C] |
215 |
by (cases "F = bot") simp_all |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
216 |
|
61245 | 217 |
lemma le_Limsup: |
218 |
assumes F: "F \<noteq> bot" and x: "\<forall>\<^sub>F x in F. l \<le> f x" |
|
219 |
shows "l \<le> Limsup F f" |
|
68860
f443ec10447d
Some basic materials on filters and topology
Manuel Eberl <eberlm@in.tum.de>
parents:
66447
diff
changeset
|
220 |
using F Liminf_bounded[of l f F] Liminf_le_Limsup[of F f] order.trans x by blast |
63895 | 221 |
|
222 |
lemma Liminf_le: |
|
223 |
assumes F: "F \<noteq> bot" and x: "\<forall>\<^sub>F x in F. f x \<le> l" |
|
224 |
shows "Liminf F f \<le> l" |
|
225 |
using F Liminf_le_Limsup Limsup_bounded order.trans x by blast |
|
61245 | 226 |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
227 |
lemma le_Liminf_iff: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
228 |
fixes X :: "_ \<Rightarrow> _ :: complete_linorder" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
229 |
shows "C \<le> Liminf F X \<longleftrightarrow> (\<forall>y<C. eventually (\<lambda>x. y < X x) F)" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
230 |
proof - |
61730 | 231 |
have "eventually (\<lambda>x. y < X x) F" |
69313 | 232 |
if "eventually P F" "y < Inf (X ` (Collect P))" for y P |
61810 | 233 |
using that by (auto elim!: eventually_mono dest: less_INF_D) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
234 |
moreover |
69313 | 235 |
have "\<exists>P. eventually P F \<and> y < Inf (X ` (Collect P))" |
61730 | 236 |
if "y < C" and y: "\<forall>y<C. eventually (\<lambda>x. y < X x) F" for y P |
237 |
proof (cases "\<exists>z. y < z \<and> z < C") |
|
238 |
case True |
|
239 |
then obtain z where z: "y < z \<and> z < C" .. |
|
69313 | 240 |
moreover from z have "z \<le> Inf (X ` {x. z < X x})" |
61730 | 241 |
by (auto intro!: INF_greatest) |
242 |
ultimately show ?thesis |
|
243 |
using y by (intro exI[of _ "\<lambda>x. z < X x"]) auto |
|
244 |
next |
|
245 |
case False |
|
69313 | 246 |
then have "C \<le> Inf (X ` {x. y < X x})" |
61730 | 247 |
by (intro INF_greatest) auto |
248 |
with \<open>y < C\<close> show ?thesis |
|
249 |
using y by (intro exI[of _ "\<lambda>x. y < X x"]) auto |
|
250 |
qed |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
251 |
ultimately show ?thesis |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
252 |
unfolding Liminf_def le_SUP_iff by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
253 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
254 |
|
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
255 |
lemma Limsup_le_iff: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
256 |
fixes X :: "_ \<Rightarrow> _ :: complete_linorder" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
257 |
shows "C \<ge> Limsup F X \<longleftrightarrow> (\<forall>y>C. eventually (\<lambda>x. y > X x) F)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
258 |
proof - |
69313 | 259 |
{ fix y P assume "eventually P F" "y > Sup (X ` (Collect P))" |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
260 |
then have "eventually (\<lambda>x. y > X x) F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
261 |
by (auto elim!: eventually_mono dest: SUP_lessD) } |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
262 |
moreover |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
263 |
{ fix y P assume "y > C" and y: "\<forall>y>C. eventually (\<lambda>x. y > X x) F" |
69313 | 264 |
have "\<exists>P. eventually P F \<and> y > Sup (X ` (Collect P))" |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
265 |
proof (cases "\<exists>z. C < z \<and> z < y") |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
266 |
case True |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
267 |
then obtain z where z: "C < z \<and> z < y" .. |
69313 | 268 |
moreover from z have "z \<ge> Sup (X ` {x. X x < z})" |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
269 |
by (auto intro!: SUP_least) |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
270 |
ultimately show ?thesis |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
271 |
using y by (intro exI[of _ "\<lambda>x. z > X x"]) auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
272 |
next |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
273 |
case False |
69313 | 274 |
then have "C \<ge> Sup (X ` {x. X x < y})" |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
275 |
by (intro SUP_least) (auto simp: not_less) |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
276 |
with \<open>y > C\<close> show ?thesis |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
277 |
using y by (intro exI[of _ "\<lambda>x. y > X x"]) auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
278 |
qed } |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
279 |
ultimately show ?thesis |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
280 |
unfolding Limsup_def INF_le_iff by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
281 |
qed |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
282 |
|
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
283 |
lemma less_LiminfD: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
284 |
"y < Liminf F (f :: _ \<Rightarrow> 'a :: complete_linorder) \<Longrightarrow> eventually (\<lambda>x. f x > y) F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
285 |
using le_Liminf_iff[of "Liminf F f" F f] by simp |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
286 |
|
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
287 |
lemma Limsup_lessD: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
288 |
"y > Limsup F (f :: _ \<Rightarrow> 'a :: complete_linorder) \<Longrightarrow> eventually (\<lambda>x. f x < y) F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
289 |
using Limsup_le_iff[of F f "Limsup F f"] by simp |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
290 |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
291 |
lemma lim_imp_Liminf: |
61730 | 292 |
fixes f :: "'a \<Rightarrow> _ :: {complete_linorder,linorder_topology}" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
293 |
assumes ntriv: "\<not> trivial_limit F" |
61973 | 294 |
assumes lim: "(f \<longlongrightarrow> f0) F" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
295 |
shows "Liminf F f = f0" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
296 |
proof (intro Liminf_eqI) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
297 |
fix P assume P: "eventually P F" |
69313 | 298 |
then have "eventually (\<lambda>x. Inf (f ` (Collect P)) \<le> f x) F" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
299 |
by eventually_elim (auto intro!: INF_lower) |
69313 | 300 |
then show "Inf (f ` (Collect P)) \<le> f0" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
301 |
by (rule tendsto_le[OF ntriv lim tendsto_const]) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
302 |
next |
69313 | 303 |
fix y assume upper: "\<And>P. eventually P F \<Longrightarrow> Inf (f ` (Collect P)) \<le> y" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
304 |
show "f0 \<le> y" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
305 |
proof cases |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
306 |
assume "\<exists>z. y < z \<and> z < f0" |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53216
diff
changeset
|
307 |
then obtain z where "y < z \<and> z < f0" .. |
69313 | 308 |
moreover have "z \<le> Inf (f ` {x. z < f x})" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
309 |
by (rule INF_greatest) simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
310 |
ultimately show ?thesis |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
311 |
using lim[THEN topological_tendstoD, THEN upper, of "{z <..}"] by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
312 |
next |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
313 |
assume discrete: "\<not> (\<exists>z. y < z \<and> z < f0)" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
314 |
show ?thesis |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
315 |
proof (rule classical) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
316 |
assume "\<not> f0 \<le> y" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
317 |
then have "eventually (\<lambda>x. y < f x) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
318 |
using lim[THEN topological_tendstoD, of "{y <..}"] by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
319 |
then have "eventually (\<lambda>x. f0 \<le> f x) F" |
61810 | 320 |
using discrete by (auto elim!: eventually_mono) |
69313 | 321 |
then have "Inf (f ` {x. f0 \<le> f x}) \<le> y" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
322 |
by (rule upper) |
69313 | 323 |
moreover have "f0 \<le> Inf (f ` {x. f0 \<le> f x})" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
324 |
by (intro INF_greatest) simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
325 |
ultimately show "f0 \<le> y" by simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
326 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
327 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
328 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
329 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
330 |
lemma lim_imp_Limsup: |
61730 | 331 |
fixes f :: "'a \<Rightarrow> _ :: {complete_linorder,linorder_topology}" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
332 |
assumes ntriv: "\<not> trivial_limit F" |
61973 | 333 |
assumes lim: "(f \<longlongrightarrow> f0) F" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
334 |
shows "Limsup F f = f0" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
335 |
proof (intro Limsup_eqI) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
336 |
fix P assume P: "eventually P F" |
69313 | 337 |
then have "eventually (\<lambda>x. f x \<le> Sup (f ` (Collect P))) F" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
338 |
by eventually_elim (auto intro!: SUP_upper) |
69313 | 339 |
then show "f0 \<le> Sup (f ` (Collect P))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
340 |
by (rule tendsto_le[OF ntriv tendsto_const lim]) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
341 |
next |
69313 | 342 |
fix y assume lower: "\<And>P. eventually P F \<Longrightarrow> y \<le> Sup (f ` (Collect P))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
343 |
show "y \<le> f0" |
53381 | 344 |
proof (cases "\<exists>z. f0 < z \<and> z < y") |
345 |
case True |
|
346 |
then obtain z where "f0 < z \<and> z < y" .. |
|
69313 | 347 |
moreover have "Sup (f ` {x. f x < z}) \<le> z" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
348 |
by (rule SUP_least) simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
349 |
ultimately show ?thesis |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
350 |
using lim[THEN topological_tendstoD, THEN lower, of "{..< z}"] by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
351 |
next |
53381 | 352 |
case False |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
353 |
show ?thesis |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
354 |
proof (rule classical) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
355 |
assume "\<not> y \<le> f0" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
356 |
then have "eventually (\<lambda>x. f x < y) F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
357 |
using lim[THEN topological_tendstoD, of "{..< y}"] by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
358 |
then have "eventually (\<lambda>x. f x \<le> f0) F" |
61810 | 359 |
using False by (auto elim!: eventually_mono simp: not_less) |
69313 | 360 |
then have "y \<le> Sup (f ` {x. f x \<le> f0})" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
361 |
by (rule lower) |
69313 | 362 |
moreover have "Sup (f ` {x. f x \<le> f0}) \<le> f0" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
363 |
by (intro SUP_least) simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
364 |
ultimately show "y \<le> f0" by simp |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
365 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
366 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
367 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
368 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
369 |
lemma Liminf_eq_Limsup: |
61730 | 370 |
fixes f0 :: "'a :: {complete_linorder,linorder_topology}" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
371 |
assumes ntriv: "\<not> trivial_limit F" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
372 |
and lim: "Liminf F f = f0" "Limsup F f = f0" |
61973 | 373 |
shows "(f \<longlongrightarrow> f0) F" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
374 |
proof (rule order_tendstoI) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
375 |
fix a assume "f0 < a" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
376 |
with assms have "Limsup F f < a" by simp |
69313 | 377 |
then obtain P where "eventually P F" "Sup (f ` (Collect P)) < a" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
378 |
unfolding Limsup_def INF_less_iff by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
379 |
then show "eventually (\<lambda>x. f x < a) F" |
61810 | 380 |
by (auto elim!: eventually_mono dest: SUP_lessD) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
381 |
next |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
382 |
fix a assume "a < f0" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
383 |
with assms have "a < Liminf F f" by simp |
69313 | 384 |
then obtain P where "eventually P F" "a < Inf (f ` (Collect P))" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
385 |
unfolding Liminf_def less_SUP_iff by auto |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
386 |
then show "eventually (\<lambda>x. a < f x) F" |
61810 | 387 |
by (auto elim!: eventually_mono dest: less_INF_D) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
388 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
389 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
390 |
lemma tendsto_iff_Liminf_eq_Limsup: |
61730 | 391 |
fixes f0 :: "'a :: {complete_linorder,linorder_topology}" |
61973 | 392 |
shows "\<not> trivial_limit F \<Longrightarrow> (f \<longlongrightarrow> f0) F \<longleftrightarrow> (Liminf F f = f0 \<and> Limsup F f = f0)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
393 |
by (metis Liminf_eq_Limsup lim_imp_Limsup lim_imp_Liminf) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
394 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
395 |
lemma liminf_subseq_mono: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
396 |
fixes X :: "nat \<Rightarrow> 'a :: complete_linorder" |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
397 |
assumes "strict_mono r" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
398 |
shows "liminf X \<le> liminf (X \<circ> r) " |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
399 |
proof- |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
400 |
have "\<And>n. (INF m\<in>{n..}. X m) \<le> (INF m\<in>{n..}. (X \<circ> r) m)" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
401 |
proof (safe intro!: INF_mono) |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
402 |
fix n m :: nat assume "n \<le> m" then show "\<exists>ma\<in>{n..}. X ma \<le> (X \<circ> r) m" |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
403 |
using seq_suble[OF \<open>strict_mono r\<close>, of m] by (intro bexI[of _ "r m"]) auto |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
404 |
qed |
56212
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents:
54261
diff
changeset
|
405 |
then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUP_INF comp_def) |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
406 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
407 |
|
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
408 |
lemma limsup_subseq_mono: |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
409 |
fixes X :: "nat \<Rightarrow> 'a :: complete_linorder" |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
410 |
assumes "strict_mono r" |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
411 |
shows "limsup (X \<circ> r) \<le> limsup X" |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
412 |
proof- |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
413 |
have "(SUP m\<in>{n..}. (X \<circ> r) m) \<le> (SUP m\<in>{n..}. X m)" for n |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
414 |
proof (safe intro!: SUP_mono) |
61730 | 415 |
fix m :: nat |
416 |
assume "n \<le> m" |
|
417 |
then show "\<exists>ma\<in>{n..}. (X \<circ> r) m \<le> X ma" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
418 |
using seq_suble[OF \<open>strict_mono r\<close>, of m] by (intro bexI[of _ "r m"]) auto |
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
419 |
qed |
61730 | 420 |
then show ?thesis |
421 |
by (auto intro!: INF_mono simp: limsup_INF_SUP comp_def) |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
422 |
qed |
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
423 |
|
61730 | 424 |
lemma continuous_on_imp_continuous_within: |
425 |
"continuous_on s f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> x \<in> s \<Longrightarrow> continuous (at x within t) f" |
|
426 |
unfolding continuous_on_eq_continuous_within |
|
427 |
by (auto simp: continuous_within intro: tendsto_within_subset) |
|
61245 | 428 |
|
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
429 |
lemma Liminf_compose_continuous_mono: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
430 |
fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
431 |
assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
432 |
shows "Liminf F (\<lambda>n. f (g n)) = f (Liminf F g)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
433 |
proof - |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
434 |
{ fix P assume "eventually P F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
435 |
have "\<exists>x. P x" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
436 |
proof (rule ccontr) |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
437 |
assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
438 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
439 |
with \<open>eventually P F\<close> F show False |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
440 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
441 |
qed } |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
442 |
note * = this |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
443 |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
444 |
have "f (SUP P\<in>{P. eventually P F}. Inf (g ` Collect P)) = |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
445 |
Sup (f ` (\<lambda>P. Inf (g ` Collect P)) ` {P. eventually P F})" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
446 |
using am continuous_on_imp_continuous_within [OF c] |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
447 |
by (rule continuous_at_Sup_mono) (auto intro: eventually_True) |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
448 |
then have "f (Liminf F g) = (SUP P \<in> {P. eventually P F}. f (Inf (g ` Collect P)))" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
449 |
by (simp add: Liminf_def image_comp) |
69313 | 450 |
also have "\<dots> = (SUP P \<in> {P. eventually P F}. Inf (f ` (g ` Collect P)))" |
69661 | 451 |
using * continuous_at_Inf_mono [OF am continuous_on_imp_continuous_within [OF c]] |
452 |
by auto |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
453 |
finally show ?thesis by (auto simp: Liminf_def image_comp) |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
454 |
qed |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
455 |
|
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
456 |
lemma Limsup_compose_continuous_mono: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
457 |
fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
458 |
assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
459 |
shows "Limsup F (\<lambda>n. f (g n)) = f (Limsup F g)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
460 |
proof - |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
461 |
{ fix P assume "eventually P F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
462 |
have "\<exists>x. P x" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
463 |
proof (rule ccontr) |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
464 |
assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
465 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
466 |
with \<open>eventually P F\<close> F show False |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
467 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
468 |
qed } |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
469 |
note * = this |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
470 |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
471 |
have "f (INF P\<in>{P. eventually P F}. Sup (g ` Collect P)) = |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
472 |
Inf (f ` (\<lambda>P. Sup (g ` Collect P)) ` {P. eventually P F})" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
473 |
using am continuous_on_imp_continuous_within [OF c] |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
474 |
by (rule continuous_at_Inf_mono) (auto intro: eventually_True) |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
475 |
then have "f (Limsup F g) = (INF P \<in> {P. eventually P F}. f (Sup (g ` Collect P)))" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
476 |
by (simp add: Limsup_def image_comp) |
69313 | 477 |
also have "\<dots> = (INF P \<in> {P. eventually P F}. Sup (f ` (g ` Collect P)))" |
69661 | 478 |
using * continuous_at_Sup_mono [OF am continuous_on_imp_continuous_within [OF c]] |
479 |
by auto |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
480 |
finally show ?thesis by (auto simp: Limsup_def image_comp) |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
481 |
qed |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
482 |
|
61245 | 483 |
lemma Liminf_compose_continuous_antimono: |
61730 | 484 |
fixes f :: "'a::{complete_linorder,linorder_topology} \<Rightarrow> 'b::{complete_linorder,linorder_topology}" |
485 |
assumes c: "continuous_on UNIV f" |
|
486 |
and am: "antimono f" |
|
487 |
and F: "F \<noteq> bot" |
|
61245 | 488 |
shows "Liminf F (\<lambda>n. f (g n)) = f (Limsup F g)" |
489 |
proof - |
|
61730 | 490 |
have *: "\<exists>x. P x" if "eventually P F" for P |
491 |
proof (rule ccontr) |
|
492 |
assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)" |
|
493 |
by auto |
|
494 |
with \<open>eventually P F\<close> F show False |
|
495 |
by auto |
|
496 |
qed |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
497 |
|
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
498 |
have "f (INF P\<in>{P. eventually P F}. Sup (g ` Collect P)) = |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
499 |
Sup (f ` (\<lambda>P. Sup (g ` Collect P)) ` {P. eventually P F})" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
500 |
using am continuous_on_imp_continuous_within [OF c] |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
501 |
by (rule continuous_at_Inf_antimono) (auto intro: eventually_True) |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
502 |
then have "f (Limsup F g) = (SUP P \<in> {P. eventually P F}. f (Sup (g ` Collect P)))" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
503 |
by (simp add: Limsup_def image_comp) |
69313 | 504 |
also have "\<dots> = (SUP P \<in> {P. eventually P F}. Inf (f ` (g ` Collect P)))" |
69661 | 505 |
using * continuous_at_Sup_antimono [OF am continuous_on_imp_continuous_within [OF c]] |
506 |
by auto |
|
61245 | 507 |
finally show ?thesis |
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
508 |
by (auto simp: Liminf_def image_comp) |
61245 | 509 |
qed |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
510 |
|
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
511 |
lemma Limsup_compose_continuous_antimono: |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
512 |
fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
513 |
assumes c: "continuous_on UNIV f" and am: "antimono f" and F: "F \<noteq> bot" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
514 |
shows "Limsup F (\<lambda>n. f (g n)) = f (Liminf F g)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
515 |
proof - |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
516 |
{ fix P assume "eventually P F" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
517 |
have "\<exists>x. P x" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
518 |
proof (rule ccontr) |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
519 |
assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)" |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
520 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
521 |
with \<open>eventually P F\<close> F show False |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
522 |
by auto |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
523 |
qed } |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
524 |
note * = this |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
525 |
|
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
526 |
have "f (SUP P\<in>{P. eventually P F}. Inf (g ` Collect P)) = |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
527 |
Inf (f ` (\<lambda>P. Inf (g ` Collect P)) ` {P. eventually P F})" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
528 |
using am continuous_on_imp_continuous_within [OF c] |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
529 |
by (rule continuous_at_Sup_antimono) (auto intro: eventually_True) |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
530 |
then have "f (Liminf F g) = (INF P \<in> {P. eventually P F}. f (Inf (g ` Collect P)))" |
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
531 |
by (simp add: Liminf_def image_comp) |
69313 | 532 |
also have "\<dots> = (INF P \<in> {P. eventually P F}. Sup (f ` (g ` Collect P)))" |
69661 | 533 |
using * continuous_at_Inf_antimono [OF am continuous_on_imp_continuous_within [OF c]] |
534 |
by auto |
|
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
535 |
finally show ?thesis |
69861
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents:
69661
diff
changeset
|
536 |
by (auto simp: Limsup_def image_comp) |
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
537 |
qed |
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
538 |
|
63895 | 539 |
lemma Liminf_filtermap_le: "Liminf (filtermap f F) g \<le> Liminf F (\<lambda>x. g (f x))" |
540 |
apply (cases "F = bot", simp) |
|
541 |
by (subst Liminf_def) |
|
542 |
(auto simp add: INF_lower Liminf_bounded eventually_filtermap eventually_mono intro!: SUP_least) |
|
543 |
||
544 |
lemma Limsup_filtermap_ge: "Limsup (filtermap f F) g \<ge> Limsup F (\<lambda>x. g (f x))" |
|
545 |
apply (cases "F = bot", simp) |
|
546 |
by (subst Limsup_def) |
|
547 |
(auto simp add: SUP_upper Limsup_bounded eventually_filtermap eventually_mono intro!: INF_greatest) |
|
548 |
||
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
549 |
lemma Liminf_least: "(\<And>P. eventually P F \<Longrightarrow> (INF x\<in>Collect P. f x) \<le> x) \<Longrightarrow> Liminf F f \<le> x" |
63895 | 550 |
by (auto intro!: SUP_least simp: Liminf_def) |
551 |
||
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
68860
diff
changeset
|
552 |
lemma Limsup_greatest: "(\<And>P. eventually P F \<Longrightarrow> x \<le> (SUP x\<in>Collect P. f x)) \<Longrightarrow> Limsup F f \<ge> x" |
63895 | 553 |
by (auto intro!: INF_greatest simp: Limsup_def) |
554 |
||
555 |
lemma Liminf_filtermap_ge: "inj f \<Longrightarrow> Liminf (filtermap f F) g \<ge> Liminf F (\<lambda>x. g (f x))" |
|
556 |
apply (cases "F = bot", simp) |
|
557 |
apply (rule Liminf_least) |
|
558 |
subgoal for P |
|
559 |
by (auto simp: eventually_filtermap the_inv_f_f |
|
560 |
intro!: Liminf_bounded INF_lower2 eventually_mono[of P]) |
|
561 |
done |
|
562 |
||
563 |
lemma Limsup_filtermap_le: "inj f \<Longrightarrow> Limsup (filtermap f F) g \<le> Limsup F (\<lambda>x. g (f x))" |
|
564 |
apply (cases "F = bot", simp) |
|
565 |
apply (rule Limsup_greatest) |
|
566 |
subgoal for P |
|
567 |
by (auto simp: eventually_filtermap the_inv_f_f |
|
568 |
intro!: Limsup_bounded SUP_upper2 eventually_mono[of P]) |
|
569 |
done |
|
570 |
||
571 |
lemma Liminf_filtermap_eq: "inj f \<Longrightarrow> Liminf (filtermap f F) g = Liminf F (\<lambda>x. g (f x))" |
|
572 |
using Liminf_filtermap_le[of f F g] Liminf_filtermap_ge[of f F g] |
|
573 |
by simp |
|
574 |
||
575 |
lemma Limsup_filtermap_eq: "inj f \<Longrightarrow> Limsup (filtermap f F) g = Limsup F (\<lambda>x. g (f x))" |
|
576 |
using Limsup_filtermap_le[of f F g] Limsup_filtermap_ge[of F g f] |
|
577 |
by simp |
|
578 |
||
62049
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
61973
diff
changeset
|
579 |
|
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
580 |
subsection \<open>More Limits\<close> |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
581 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
582 |
lemma convergent_limsup_cl: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
583 |
fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
584 |
shows "convergent X \<Longrightarrow> limsup X = lim X" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
585 |
by (auto simp: convergent_def limI lim_imp_Limsup) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
586 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
587 |
lemma convergent_liminf_cl: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
588 |
fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
589 |
shows "convergent X \<Longrightarrow> liminf X = lim X" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
590 |
by (auto simp: convergent_def limI lim_imp_Liminf) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
591 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
592 |
lemma lim_increasing_cl: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
593 |
assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<ge> f m" |
61969 | 594 |
obtains l where "f \<longlonglongrightarrow> (l::'a::{complete_linorder,linorder_topology})" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
595 |
proof |
61969 | 596 |
show "f \<longlonglongrightarrow> (SUP n. f n)" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
597 |
using assms |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
598 |
by (intro increasing_tendsto) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
599 |
(auto simp: SUP_upper eventually_sequentially less_SUP_iff intro: less_le_trans) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
600 |
qed |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
601 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
602 |
lemma lim_decreasing_cl: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
603 |
assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<le> f m" |
61969 | 604 |
obtains l where "f \<longlonglongrightarrow> (l::'a::{complete_linorder,linorder_topology})" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
605 |
proof |
61969 | 606 |
show "f \<longlonglongrightarrow> (INF n. f n)" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
607 |
using assms |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
608 |
by (intro decreasing_tendsto) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
609 |
(auto simp: INF_lower eventually_sequentially INF_less_iff intro: le_less_trans) |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
610 |
qed |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
611 |
|
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
612 |
lemma compact_complete_linorder: |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
613 |
fixes X :: "nat \<Rightarrow> 'a::{complete_linorder,linorder_topology}" |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
614 |
shows "\<exists>l r. strict_mono r \<and> (X \<circ> r) \<longlonglongrightarrow> l" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
615 |
proof - |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
616 |
obtain r where "strict_mono r" and mono: "monoseq (X \<circ> r)" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
617 |
using seq_monosub[of X] |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
618 |
unfolding comp_def |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
619 |
by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
620 |
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)" |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
621 |
by (auto simp add: monoseq_def) |
61969 | 622 |
then obtain l where "(X \<circ> r) \<longlonglongrightarrow> l" |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
623 |
using lim_increasing_cl[of "X \<circ> r"] lim_decreasing_cl[of "X \<circ> r"] |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
624 |
by auto |
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
625 |
then show ?thesis |
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
63895
diff
changeset
|
626 |
using \<open>strict_mono r\<close> by auto |
61880
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents:
61810
diff
changeset
|
627 |
qed |
61245 | 628 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
629 |
lemma tendsto_Limsup: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
630 |
fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
631 |
shows "F \<noteq> bot \<Longrightarrow> Limsup F f = Liminf F f \<Longrightarrow> (f \<longlongrightarrow> Limsup F f) F" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
632 |
by (subst tendsto_iff_Liminf_eq_Limsup) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
633 |
|
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
634 |
lemma tendsto_Liminf: |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
635 |
fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
636 |
shows "F \<noteq> bot \<Longrightarrow> Limsup F f = Liminf F f \<Longrightarrow> (f \<longlongrightarrow> Liminf F f) F" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
637 |
by (subst tendsto_iff_Liminf_eq_Limsup) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62624
diff
changeset
|
638 |
|
51340
5e6296afe08d
move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff
changeset
|
639 |
end |