src/HOL/Library/Liminf_Limsup.thy
author nipkow
Wed, 13 Feb 2019 07:48:42 +0100
changeset 69801 a99a0f5474c5
parent 69661 a03a63b81f44
child 69861 62e47f06d22c
permissions -rw-r--r--
too agressive
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
738598beeb26 tuned imports;
wenzelm
parents: 51340
diff changeset
     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
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    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
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51542
diff changeset
    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
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    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
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 51542
diff changeset
    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
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61245
diff changeset
    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
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
    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
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
    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
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    92
  "(\<And>P. eventually P F \<Longrightarrow> Inf (f ` (Collect P)) \<le> x) \<Longrightarrow>
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    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
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    97
  "(\<And>P. eventually P F \<Longrightarrow> x \<le> Sup (f ` (Collect P))) \<Longrightarrow>
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
    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
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   109
lemma Limsup_const:
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   110
  assumes ntriv: "\<not> trivial_limit F"
54261
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
   111
  shows "Limsup F (\<lambda>x. c) = c"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   112
proof -
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   113
  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
   114
  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
   115
    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
   116
  then show ?thesis
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   117
    apply (auto simp add: Limsup_def)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   118
    apply (rule INF_const)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   119
    apply auto
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   120
    using eventually_True apply blast
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   121
    done
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   122
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   123
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   124
lemma Liminf_const:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   125
  assumes ntriv: "\<not> trivial_limit F"
54261
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
   126
  shows "Liminf F (\<lambda>x. c) = c"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   127
proof -
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   128
  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
   129
  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
   130
    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
   131
  then show ?thesis
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   132
    apply (auto simp add: Liminf_def)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   133
    apply (rule SUP_const)
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   134
    apply auto
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   135
    using eventually_True apply blast
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   136
    done
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   137
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   138
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   139
lemma Liminf_mono:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   140
  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
   141
  shows "Liminf F f \<le> Liminf F g"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   142
  unfolding Liminf_def
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   143
proof (safe intro!: SUP_mono)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   144
  fix P assume "eventually P F"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   145
  with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   146
  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
   147
    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
   148
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   149
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   150
lemma Liminf_eq:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   151
  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
   152
  shows "Liminf F f = Liminf F g"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   153
  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
   154
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   155
lemma Limsup_mono:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   156
  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
   157
  shows "Limsup F f \<le> Limsup F g"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   158
  unfolding Limsup_def
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   159
proof (safe intro!: INF_mono)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   160
  fix P assume "eventually P F"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   161
  with ev have "eventually (\<lambda>x. f x \<le> g x \<and> P x) F" (is "eventually ?Q F") by (rule eventually_conj)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   162
  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
   163
    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
   164
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   165
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   166
lemma Limsup_eq:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   167
  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
   168
  shows "Limsup net f = Limsup net g"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   169
  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
   170
63895
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   171
lemma Liminf_bot[simp]: "Liminf bot f = top"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   172
  unfolding Liminf_def top_unique[symmetric]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   173
  by (rule SUP_upper2[where i="\<lambda>x. False"]) simp_all
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   174
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   175
lemma Limsup_bot[simp]: "Limsup bot f = bot"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   176
  unfolding Limsup_def bot_unique[symmetric]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   177
  by (rule INF_lower2[where i="\<lambda>x. False"]) simp_all
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   178
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   179
lemma Liminf_le_Limsup:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   180
  assumes ntriv: "\<not> trivial_limit F"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   181
  shows "Liminf F f \<le> Limsup F f"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   182
  unfolding Limsup_def Liminf_def
54261
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
   183
  apply (rule SUP_least)
89991ef58448 restrict Limsup and Liminf to complete lattices
hoelzl
parents: 54257
diff changeset
   184
  apply (rule INF_greatest)
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   185
proof safe
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   186
  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
   187
  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
   188
  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
   189
    using ntriv by (auto simp add: eventually_False)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   190
  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
   191
    by (rule INF_mono) auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   192
  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
   193
    using not_False by (intro INF_le_SUP) auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   194
  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
   195
    by (rule SUP_mono) auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   196
  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
   197
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   198
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   199
lemma Liminf_bounded:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   200
  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
   201
  shows "C \<le> Liminf F X"
63895
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   202
  using Liminf_mono[OF le] Liminf_const[of F C]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   203
  by (cases "F = bot") simp_all
51340
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 Limsup_bounded:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   206
  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
   207
  shows "Limsup F X \<le> C"
63895
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   208
  using Limsup_mono[OF le] Limsup_const[of F C]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   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
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   211
lemma le_Limsup:
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   212
  assumes F: "F \<noteq> bot" and x: "\<forall>\<^sub>F x in F. l \<le> f x"
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   213
  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
   214
  using F Liminf_bounded[of l f F] Liminf_le_Limsup[of F f] order.trans x by blast
63895
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   215
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   216
lemma Liminf_le:
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   217
  assumes F: "F \<noteq> bot" and x: "\<forall>\<^sub>F x in F. f x \<le> l"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   218
  shows "Liminf F f \<le> l"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   219
  using F Liminf_le_Limsup Limsup_bounded order.trans x by blast
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   220
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   221
lemma le_Liminf_iff:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   222
  fixes X :: "_ \<Rightarrow> _ :: complete_linorder"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   223
  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
   224
proof -
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   225
  have "eventually (\<lambda>x. y < X x) F"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   226
    if "eventually P F" "y < Inf (X ` (Collect P))" for y P
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   227
    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
   228
  moreover
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   229
  have "\<exists>P. eventually P F \<and> y < Inf (X ` (Collect P))"
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   230
    if "y < C" and y: "\<forall>y<C. eventually (\<lambda>x. y < X x) F" for y P
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   231
  proof (cases "\<exists>z. y < z \<and> z < C")
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   232
    case True
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   233
    then obtain z where z: "y < z \<and> z < C" ..
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   234
    moreover from z have "z \<le> Inf (X ` {x. z < X x})"
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   235
      by (auto intro!: INF_greatest)
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   236
    ultimately show ?thesis
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   237
      using y by (intro exI[of _ "\<lambda>x. z < X x"]) auto
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   238
  next
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   239
    case False
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   240
    then have "C \<le> Inf (X ` {x. y < X x})"
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   241
      by (intro INF_greatest) auto
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   242
    with \<open>y < C\<close> show ?thesis
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   243
      using y by (intro exI[of _ "\<lambda>x. y < X x"]) auto
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   244
  qed
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   245
  ultimately show ?thesis
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   246
    unfolding Liminf_def le_SUP_iff by auto
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   247
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   248
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   249
lemma Limsup_le_iff:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   250
  fixes X :: "_ \<Rightarrow> _ :: complete_linorder"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   251
  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
   252
proof -
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   253
  { 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
   254
    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
   255
      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
   256
  moreover
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   257
  { fix y P assume "y > C" and y: "\<forall>y>C. eventually (\<lambda>x. y > X x) F"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   258
    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
   259
    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
   260
      case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   261
      then obtain z where z: "C < z \<and> z < y" ..
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   262
      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
   263
        by (auto intro!: SUP_least)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   264
      ultimately show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   265
        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
   266
    next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   267
      case False
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   268
      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
   269
        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
   270
      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
   271
        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
   272
    qed }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   273
  ultimately show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   274
    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
   275
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   276
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   277
lemma less_LiminfD:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   278
  "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
   279
  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
   280
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   281
lemma Limsup_lessD:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   282
  "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
   283
  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
   284
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   285
lemma lim_imp_Liminf:
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   286
  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
   287
  assumes ntriv: "\<not> trivial_limit F"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   288
  assumes lim: "(f \<longlongrightarrow> f0) F"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   289
  shows "Liminf F f = f0"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   290
proof (intro Liminf_eqI)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   291
  fix P assume P: "eventually P F"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   292
  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
   293
    by eventually_elim (auto intro!: INF_lower)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   294
  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
   295
    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
   296
next
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   297
  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
   298
  show "f0 \<le> y"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   299
  proof cases
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   300
    assume "\<exists>z. y < z \<and> z < f0"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   301
    then obtain z where "y < z \<and> z < f0" ..
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   302
    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
   303
      by (rule INF_greatest) simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   304
    ultimately show ?thesis
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   305
      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
   306
  next
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   307
    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
   308
    show ?thesis
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   309
    proof (rule classical)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   310
      assume "\<not> f0 \<le> y"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   311
      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
   312
        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
   313
      then have "eventually (\<lambda>x. f0 \<le> f x) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   314
        using discrete by (auto elim!: eventually_mono)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   315
      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
   316
        by (rule upper)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   317
      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
   318
        by (intro INF_greatest) simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   319
      ultimately show "f0 \<le> y" by simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   320
    qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   321
  qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   322
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   323
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   324
lemma lim_imp_Limsup:
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   325
  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
   326
  assumes ntriv: "\<not> trivial_limit F"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   327
  assumes lim: "(f \<longlongrightarrow> f0) F"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   328
  shows "Limsup F f = f0"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   329
proof (intro Limsup_eqI)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   330
  fix P assume P: "eventually P F"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   331
  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
   332
    by eventually_elim (auto intro!: SUP_upper)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   333
  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
   334
    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
   335
next
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   336
  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
   337
  show "y \<le> f0"
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53374
diff changeset
   338
  proof (cases "\<exists>z. f0 < z \<and> z < y")
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53374
diff changeset
   339
    case True
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53374
diff changeset
   340
    then obtain z where "f0 < z \<and> z < y" ..
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   341
    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
   342
      by (rule SUP_least) simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   343
    ultimately show ?thesis
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   344
      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
   345
  next
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53374
diff changeset
   346
    case False
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   347
    show ?thesis
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   348
    proof (rule classical)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   349
      assume "\<not> y \<le> f0"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   350
      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
   351
        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
   352
      then have "eventually (\<lambda>x. f x \<le> f0) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   353
        using False by (auto elim!: eventually_mono simp: not_less)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   354
      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
   355
        by (rule lower)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   356
      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
   357
        by (intro SUP_least) simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   358
      ultimately show "y \<le> f0" by simp
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   359
    qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   360
  qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   361
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   362
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   363
lemma Liminf_eq_Limsup:
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   364
  fixes f0 :: "'a :: {complete_linorder,linorder_topology}"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   365
  assumes ntriv: "\<not> trivial_limit F"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   366
    and lim: "Liminf F f = f0" "Limsup F f = f0"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   367
  shows "(f \<longlongrightarrow> f0) F"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   368
proof (rule order_tendstoI)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   369
  fix a assume "f0 < a"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   370
  with assms have "Limsup F f < a" by simp
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   371
  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
   372
    unfolding Limsup_def INF_less_iff by auto
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   373
  then show "eventually (\<lambda>x. f x < a) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   374
    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
   375
next
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   376
  fix a assume "a < f0"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   377
  with assms have "a < Liminf F f" by simp
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   378
  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
   379
    unfolding Liminf_def less_SUP_iff by auto
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   380
  then show "eventually (\<lambda>x. a < f x) F"
61810
3c5040d5694a sorted out eventually_mono
paulson <lp15@cam.ac.uk>
parents: 61806
diff changeset
   381
    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
   382
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   383
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   384
lemma tendsto_iff_Liminf_eq_Limsup:
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   385
  fixes f0 :: "'a :: {complete_linorder,linorder_topology}"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   386
  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
   387
  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
   388
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   389
lemma liminf_subseq_mono:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   390
  fixes X :: "nat \<Rightarrow> 'a :: complete_linorder"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 63895
diff changeset
   391
  assumes "strict_mono r"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   392
  shows "liminf X \<le> liminf (X \<circ> r) "
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   393
proof-
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   394
  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
   395
  proof (safe intro!: INF_mono)
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   396
    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
   397
      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
   398
  qed
56212
3253aaf73a01 consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
haftmann
parents: 54261
diff changeset
   399
  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
   400
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   401
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   402
lemma limsup_subseq_mono:
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   403
  fixes X :: "nat \<Rightarrow> 'a :: complete_linorder"
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 63895
diff changeset
   404
  assumes "strict_mono r"
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   405
  shows "limsup (X \<circ> r) \<le> limsup X"
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   406
proof-
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   407
  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
   408
  proof (safe intro!: SUP_mono)
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   409
    fix m :: nat
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   410
    assume "n \<le> m"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   411
    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
   412
      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
   413
  qed
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   414
  then show ?thesis
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   415
    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
   416
qed
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   417
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   418
lemma continuous_on_imp_continuous_within:
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   419
  "continuous_on s f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> x \<in> s \<Longrightarrow> continuous (at x within t) f"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   420
  unfolding continuous_on_eq_continuous_within
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   421
  by (auto simp: continuous_within intro: tendsto_within_subset)
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   422
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   423
lemma Liminf_compose_continuous_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   424
  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
   425
  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
   426
  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
   427
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   428
  { 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
   429
    have "\<exists>x. P x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   430
    proof (rule ccontr)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   431
      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
   432
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   433
      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
   434
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   435
    qed }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   436
  note * = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   437
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   438
  have "f (Liminf F g) = (SUP P \<in> {P. eventually P F}. f (Inf (g ` Collect P)))"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62049
diff changeset
   439
    unfolding Liminf_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   440
    by (subst continuous_at_Sup_mono[OF am continuous_on_imp_continuous_within[OF c]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   441
       (auto intro: eventually_True)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   442
  also have "\<dots> = (SUP P \<in> {P. eventually P F}. Inf (f ` (g ` Collect P)))"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   443
    using * continuous_at_Inf_mono [OF am continuous_on_imp_continuous_within [OF c]]
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   444
    by auto 
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   445
  finally show ?thesis by (auto simp: Liminf_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   446
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   447
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   448
lemma Limsup_compose_continuous_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   449
  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
   450
  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
   451
  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
   452
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   453
  { 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
   454
    have "\<exists>x. P x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   455
    proof (rule ccontr)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   456
      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
   457
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   458
      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
   459
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   460
    qed }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   461
  note * = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   462
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   463
  have "f (Limsup F g) = (INF P \<in> {P. eventually P F}. f (Sup (g ` Collect P)))"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62049
diff changeset
   464
    unfolding Limsup_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   465
    by (subst continuous_at_Inf_mono[OF am continuous_on_imp_continuous_within[OF c]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   466
       (auto intro: eventually_True)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   467
  also have "\<dots> = (INF P \<in> {P. eventually P F}. Sup (f ` (g ` Collect P)))"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   468
    using * continuous_at_Sup_mono [OF am continuous_on_imp_continuous_within [OF c]]
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   469
    by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   470
  finally show ?thesis by (auto simp: Limsup_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   471
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   472
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   473
lemma Liminf_compose_continuous_antimono:
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   474
  fixes f :: "'a::{complete_linorder,linorder_topology} \<Rightarrow> 'b::{complete_linorder,linorder_topology}"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   475
  assumes c: "continuous_on UNIV f"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   476
    and am: "antimono f"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   477
    and F: "F \<noteq> bot"
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   478
  shows "Liminf F (\<lambda>n. f (g n)) = f (Limsup F g)"
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   479
proof -
61730
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   480
  have *: "\<exists>x. P x" if "eventually P F" for P
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   481
  proof (rule ccontr)
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   482
    assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   483
      by auto
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   484
    with \<open>eventually P F\<close> F show False
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   485
      by auto
2b775b888897 tuned proofs;
wenzelm
parents: 61585
diff changeset
   486
  qed
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   487
  have "f (Limsup F g) = (SUP P \<in> {P. eventually P F}. f (Sup (g ` Collect P)))"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62049
diff changeset
   488
    unfolding Limsup_def
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   489
    by (subst continuous_at_Inf_antimono[OF am continuous_on_imp_continuous_within[OF c]])
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   490
       (auto intro: eventually_True)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   491
  also have "\<dots> = (SUP P \<in> {P. eventually P F}. Inf (f ` (g ` Collect P)))"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   492
    using * continuous_at_Sup_antimono [OF am continuous_on_imp_continuous_within [OF c]]
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   493
    by auto
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   494
  finally show ?thesis
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   495
    by (auto simp: Liminf_def)
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   496
qed
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   497
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   498
lemma Limsup_compose_continuous_antimono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   499
  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
   500
  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
   501
  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
   502
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   503
  { 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
   504
    have "\<exists>x. P x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   505
    proof (rule ccontr)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   506
      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
   507
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   508
      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
   509
        by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   510
    qed }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   511
  note * = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   512
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   513
  have "f (Liminf F g) = (INF P \<in> {P. eventually P F}. f (Inf (g ` Collect P)))"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62049
diff changeset
   514
    unfolding Liminf_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   515
    by (subst continuous_at_Sup_antimono[OF am continuous_on_imp_continuous_within[OF c]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   516
       (auto intro: eventually_True)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   517
  also have "\<dots> = (INF P \<in> {P. eventually P F}. Sup (f ` (g ` Collect P)))"
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   518
    using * continuous_at_Inf_antimono [OF am continuous_on_imp_continuous_within [OF c]]
a03a63b81f44 tuned proofs
haftmann
parents: 69313
diff changeset
   519
    by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   520
  finally show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   521
    by (auto simp: Limsup_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   522
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   523
63895
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   524
lemma Liminf_filtermap_le: "Liminf (filtermap f F) g \<le> Liminf F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   525
  apply (cases "F = bot", simp)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   526
  by (subst Liminf_def)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   527
    (auto simp add: INF_lower Liminf_bounded eventually_filtermap eventually_mono intro!: SUP_least)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   528
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   529
lemma Limsup_filtermap_ge: "Limsup (filtermap f F) g \<ge> Limsup F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   530
  apply (cases "F = bot", simp)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   531
  by (subst Limsup_def)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   532
    (auto simp add: SUP_upper Limsup_bounded eventually_filtermap eventually_mono intro!: INF_greatest)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   533
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   534
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
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   535
  by (auto intro!: SUP_least simp: Liminf_def)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   536
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 68860
diff changeset
   537
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
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   538
  by (auto intro!: INF_greatest simp: Limsup_def)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   539
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   540
lemma Liminf_filtermap_ge: "inj f \<Longrightarrow> Liminf (filtermap f F) g \<ge> Liminf F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   541
  apply (cases "F = bot", simp)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   542
  apply (rule Liminf_least)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   543
  subgoal for P
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   544
    by (auto simp: eventually_filtermap the_inv_f_f
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   545
        intro!: Liminf_bounded INF_lower2 eventually_mono[of P])
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   546
  done
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   547
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   548
lemma Limsup_filtermap_le: "inj f \<Longrightarrow> Limsup (filtermap f F) g \<le> Limsup F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   549
  apply (cases "F = bot", simp)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   550
  apply (rule Limsup_greatest)
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   551
  subgoal for P
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   552
    by (auto simp: eventually_filtermap the_inv_f_f
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   553
        intro!: Limsup_bounded SUP_upper2 eventually_mono[of P])
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   554
  done
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   555
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   556
lemma Liminf_filtermap_eq: "inj f \<Longrightarrow> Liminf (filtermap f F) g = Liminf F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   557
  using Liminf_filtermap_le[of f F g] Liminf_filtermap_ge[of f F g]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   558
  by simp
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   559
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   560
lemma Limsup_filtermap_eq: "inj f \<Longrightarrow> Limsup (filtermap f F) g = Limsup F (\<lambda>x. g (f x))"
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   561
  using Limsup_filtermap_le[of f F g] Limsup_filtermap_ge[of F g f]
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   562
  by simp
9afa979137da Liminf/Limsup and filtermap
immler
parents: 62975
diff changeset
   563
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
   564
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   565
subsection \<open>More Limits\<close>
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   566
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   567
lemma convergent_limsup_cl:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   568
  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
   569
  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
   570
  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
   571
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   572
lemma convergent_liminf_cl:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   573
  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
   574
  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
   575
  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
   576
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   577
lemma lim_increasing_cl:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   578
  assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<ge> f m"
61969
e01015e49041 more symbols;
wenzelm
parents: 61880
diff changeset
   579
  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
   580
proof
61969
e01015e49041 more symbols;
wenzelm
parents: 61880
diff changeset
   581
  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
   582
    using assms
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   583
    by (intro increasing_tendsto)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   584
       (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
   585
qed
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 lim_decreasing_cl:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   588
  assumes "\<And>n m. n \<ge> m \<Longrightarrow> f n \<le> f m"
61969
e01015e49041 more symbols;
wenzelm
parents: 61880
diff changeset
   589
  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
   590
proof
61969
e01015e49041 more symbols;
wenzelm
parents: 61880
diff changeset
   591
  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
   592
    using assms
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   593
    by (intro decreasing_tendsto)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   594
       (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
   595
qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   596
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   597
lemma compact_complete_linorder:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   598
  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
   599
  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
   600
proof -
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 63895
diff changeset
   601
  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
   602
    using seq_monosub[of X]
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   603
    unfolding comp_def
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   604
    by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   605
  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
   606
    by (auto simp add: monoseq_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61880
diff changeset
   607
  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
   608
     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
   609
     by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61810
diff changeset
   610
  then show ?thesis
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 63895
diff changeset
   611
    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
   612
qed
61245
b77bf45efe21 prove Liminf_inverse_ereal
hoelzl
parents: 60500
diff changeset
   613
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   614
lemma tendsto_Limsup:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   615
  fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   616
  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
   617
  by (subst tendsto_iff_Liminf_eq_Limsup) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   618
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   619
lemma tendsto_Liminf:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   620
  fixes f :: "_ \<Rightarrow> 'a :: {complete_linorder,linorder_topology}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   621
  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
   622
  by (subst tendsto_iff_Liminf_eq_Limsup) auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62624
diff changeset
   623
51340
5e6296afe08d move Liminf / Limsup lemmas on complete_lattices to its own file
hoelzl
parents:
diff changeset
   624
end